diff --git a/stage0/src/CMakeLists.txt b/stage0/src/CMakeLists.txt index 95eece17470a..2f03e98dc4db 100644 --- a/stage0/src/CMakeLists.txt +++ b/stage0/src/CMakeLists.txt @@ -698,17 +698,17 @@ else() endif() if(NOT ${CMAKE_SYSTEM_NAME} MATCHES "Emscripten") - add_custom_target(lake_lib ALL + add_custom_target(lake_lib WORKING_DIRECTORY ${LEAN_SOURCE_DIR} DEPENDS leanshared COMMAND $(MAKE) -f ${CMAKE_BINARY_DIR}/stdlib.make Lake VERBATIM) - add_custom_target(lake_shared ALL + add_custom_target(lake_shared WORKING_DIRECTORY ${LEAN_SOURCE_DIR} DEPENDS lake_lib COMMAND $(MAKE) -f ${CMAKE_BINARY_DIR}/stdlib.make libLake_shared VERBATIM) - add_custom_target(lake ALL + add_custom_target(lake #ALL WORKING_DIRECTORY ${LEAN_SOURCE_DIR} DEPENDS lake_shared COMMAND $(MAKE) -f ${CMAKE_BINARY_DIR}/stdlib.make lake diff --git a/stage0/src/kernel/environment.cpp b/stage0/src/kernel/environment.cpp index 57e0a0358ec4..1e8514eff38f 100644 --- a/stage0/src/kernel/environment.cpp +++ b/stage0/src/kernel/environment.cpp @@ -19,16 +19,9 @@ Author: Leonardo de Moura namespace lean { extern "C" object* lean_environment_add(object*, object*); -extern "C" object* lean_mk_empty_environment(uint32, object*); extern "C" object* lean_environment_find(object*, object*); -extern "C" uint32 lean_environment_trust_level(object*); extern "C" object* lean_environment_mark_quot_init(object*); extern "C" uint8 lean_environment_quot_init(object*); -extern "C" object* lean_register_extension(object*); -extern "C" object* lean_get_extension(object*, object*); -extern "C" object* lean_set_extension(object*, object*, object*); -extern "C" object* lean_environment_set_main_module(object*, object*); -extern "C" object* lean_environment_main_module(object*); extern "C" object* lean_kernel_record_unfold (object*, object*); extern "C" object* lean_kernel_get_diag(object*); extern "C" object* lean_kernel_set_diag(object*, object*); @@ -62,14 +55,6 @@ environment scoped_diagnostics::update(environment const & env) const { return env; } -environment mk_empty_environment(uint32 trust_lvl) { - return get_io_result(lean_mk_empty_environment(trust_lvl, io_mk_world())); -} - -environment::environment(unsigned trust_lvl): - object_ref(mk_empty_environment(trust_lvl)) { -} - diagnostics environment::get_diag() const { return diagnostics(lean_kernel_get_diag(to_obj_arg())); } @@ -78,18 +63,6 @@ environment environment::set_diag(diagnostics const & diag) const { return environment(lean_kernel_set_diag(to_obj_arg(), diag.to_obj_arg())); } -void environment::set_main_module(name const & n) { - m_obj = lean_environment_set_main_module(m_obj, n.to_obj_arg()); -} - -name environment::get_main_module() const { - return name(lean_environment_main_module(to_obj_arg())); -} - -unsigned environment::trust_lvl() const { - return lean_environment_trust_level(to_obj_arg()); -} - bool environment::is_quot_initialized() const { return lean_environment_quot_init(to_obj_arg()) != 0; } @@ -297,7 +270,7 @@ environment environment::add(declaration const & d, bool check) const { } /* addDeclCore (env : Environment) (maxHeartbeats : USize) (decl : @& Declaration) - (cancelTk? : @& Option IO.CancelToken) : Except KernelException Environment + (cancelTk? : @& Option IO.CancelToken) : Except Kernel.Exception Environment */ extern "C" LEAN_EXPORT object * lean_add_decl(object * env, size_t max_heartbeat, object * decl, object * opt_cancel_tk) { @@ -321,12 +294,6 @@ void environment::for_each_constant(std::function }); } -extern "C" obj_res lean_display_stats(obj_arg env, obj_arg w); - -void environment::display_stats() const { - dec_ref(lean_display_stats(to_obj_arg(), io_mk_world())); -} - void initialize_environment() { } diff --git a/stage0/src/kernel/environment.h b/stage0/src/kernel/environment.h index d0d61a8eed9b..88ff1496995a 100644 --- a/stage0/src/kernel/environment.h +++ b/stage0/src/kernel/environment.h @@ -24,10 +24,6 @@ Author: Leonardo de Moura #endif namespace lean { -class environment_extension { -public: - virtual ~environment_extension() {} -}; /* Wrapper for `Kernel.Diagnostics` */ class diagnostics : public object_ref { @@ -41,7 +37,7 @@ class diagnostics : public object_ref { }; /* -Store `Kernel.Diagnostics` stored in environment extension in `m_diag` IF +Store `Kernel.Diagnostics` (to be stored in `Kernel.Environment.diagnostics`) in `m_diag` IF - `Kernel.Diagnostics.enable = true` - `collect = true`. This is a minor optimization. @@ -59,6 +55,7 @@ class scoped_diagnostics { diagnostics * get() const { return m_diag; } }; +/* Wrapper for `Lean.Kernel.Environment` */ class LEAN_EXPORT environment : public object_ref { friend class add_inductive_fn; @@ -76,7 +73,6 @@ class LEAN_EXPORT environment : public object_ref { environment add_quot() const; environment add_inductive(declaration const & d) const; public: - environment(unsigned trust_lvl = 0); environment(environment const & other):object_ref(other) {} environment(environment && other):object_ref(other) {} explicit environment(b_obj_arg o, bool b):object_ref(o, b) {} @@ -89,15 +85,8 @@ class LEAN_EXPORT environment : public object_ref { diagnostics get_diag() const; environment set_diag(diagnostics const & diag) const; - /** \brief Return the trust level of this environment. */ - unsigned trust_lvl() const; - bool is_quot_initialized() const; - void set_main_module(name const & n); - - name get_main_module() const; - /** \brief Return information for the constant with name \c n (if it is defined in this environment). */ optional find(name const & n) const; @@ -114,8 +103,6 @@ class LEAN_EXPORT environment : public object_ref { friend bool is_eqp(environment const & e1, environment const & e2) { return e1.raw() == e2.raw(); } - - void display_stats() const; }; void check_no_metavar_no_fvar(environment const & env, name const & n, expr const & e); diff --git a/stage0/src/kernel/kernel_exception.h b/stage0/src/kernel/kernel_exception.h index 793cf6403277..243b53635fda 100644 --- a/stage0/src/kernel/kernel_exception.h +++ b/stage0/src/kernel/kernel_exception.h @@ -152,9 +152,9 @@ class invalid_proj_exception : public kernel_exception_with_lctx { /* Helper function for interfacing C++ code with code written in Lean. It executes closure `f` which produces an object_ref of type `A` and may throw -an `kernel_exception` or `exception`. Then, convert result into `Except KernelException T` +an `kernel_exception` or `exception`. Then, convert result into `Except Kernel.Exception T` where `T` is the type of the lean objected represented by `A`. -We use the constructor `KernelException.other ` to handle C++ `exception` objects which +We use the constructor `Kernel.Exception.other ` to handle C++ `exception` objects which are not `kernel_exception`. */ template diff --git a/stage0/src/kernel/trace.cpp b/stage0/src/kernel/trace.cpp index 9cec464e5fd6..c556540612b9 100644 --- a/stage0/src/kernel/trace.cpp +++ b/stage0/src/kernel/trace.cpp @@ -8,7 +8,7 @@ Author: Leonardo de Moura #include #include "util/io.h" #include "util/option_declarations.h" -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "kernel/local_ctx.h" #include "kernel/trace.h" @@ -17,7 +17,7 @@ static name_set * g_trace_classes = nullptr; static name_map * g_trace_aliases = nullptr; MK_THREAD_LOCAL_GET_DEF(std::vector, get_enabled_trace_classes); MK_THREAD_LOCAL_GET_DEF(std::vector, get_disabled_trace_classes); -LEAN_THREAD_PTR(environment, g_env); +LEAN_THREAD_PTR(elab_environment, g_env); LEAN_THREAD_PTR(options, g_opts); void register_trace_class(name const & n, name const & decl_name) { @@ -90,7 +90,7 @@ bool is_trace_class_enabled(name const & n) { } -void scope_trace_env::init(environment * env, options * opts) { +void scope_trace_env::init(elab_environment * env, options * opts) { m_enable_sz = get_enabled_trace_classes().size(); m_disable_sz = get_disabled_trace_classes().size(); m_old_env = g_env; @@ -111,12 +111,12 @@ void scope_trace_env::init(environment * env, options * opts) { g_opts = opts; } -scope_trace_env::scope_trace_env(environment const & env, options const & o) { - init(const_cast(&env), const_cast(&o)); +scope_trace_env::scope_trace_env(elab_environment const & env, options const & o) { + init(const_cast(&env), const_cast(&o)); } scope_trace_env::~scope_trace_env() { - g_env = const_cast(m_old_env); + g_env = const_cast(m_old_env); g_opts = const_cast(m_old_opts); get_enabled_trace_classes().resize(m_enable_sz); get_disabled_trace_classes().resize(m_disable_sz); @@ -169,7 +169,7 @@ def pretty (f : Format) (w : Nat := defWidth) (indent : Nat := 0) (column := 0) */ extern "C" object * lean_format_pretty(object * f, object * w, object * i, object * c); -std::string pp_expr(environment const & env, options const & opts, local_ctx const & lctx, expr const & e) { +std::string pp_expr(elab_environment const & env, options const & opts, local_ctx const & lctx, expr const & e) { options o = opts; // o = o.update(name{"pp", "proofs"}, true); -- object_ref fmt = get_io_result(lean_pp_expr(env.to_obj_arg(), lean_mk_metavar_ctx(lean_box(0)), lctx.to_obj_arg(), o.to_obj_arg(), @@ -178,7 +178,7 @@ std::string pp_expr(environment const & env, options const & opts, local_ctx con return str.to_std_string(); } -std::string pp_expr(environment const & env, options const & opts, expr const & e) { +std::string pp_expr(elab_environment const & env, options const & opts, expr const & e) { local_ctx lctx; return pp_expr(env, opts, lctx, e); } diff --git a/stage0/src/kernel/trace.h b/stage0/src/kernel/trace.h index feac201c88e7..e42f20f90d13 100644 --- a/stage0/src/kernel/trace.h +++ b/stage0/src/kernel/trace.h @@ -7,7 +7,7 @@ Author: Leonardo de Moura #pragma once #include #include -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "util/options.h" #include "util/message_definitions.h" @@ -20,13 +20,13 @@ bool is_trace_class_enabled(name const & n); #define lean_is_trace_enabled(CName) (::lean::is_trace_enabled() && ::lean::is_trace_class_enabled(CName)) class scope_trace_env { - unsigned m_enable_sz; - unsigned m_disable_sz; - environment const * m_old_env; - options const * m_old_opts; - void init(environment * env, options * opts); + unsigned m_enable_sz; + unsigned m_disable_sz; + elab_environment const * m_old_env; + options const * m_old_opts; + void init(elab_environment * env, options * opts); public: - scope_trace_env(environment const & env, options const & opts); + scope_trace_env(elab_environment const & env, options const & opts); ~scope_trace_env(); }; diff --git a/stage0/src/kernel/type_checker.cpp b/stage0/src/kernel/type_checker.cpp index b0e6844dca24..dc4034e9ffc6 100644 --- a/stage0/src/kernel/type_checker.cpp +++ b/stage0/src/kernel/type_checker.cpp @@ -549,7 +549,7 @@ static expr * g_lean_reduce_bool = nullptr; static expr * g_lean_reduce_nat = nullptr; namespace ir { -object * run_boxed(environment const & env, options const & opts, name const & fn, unsigned n, object **args); +object * run_boxed_kernel(environment const & env, options const & opts, name const & fn, unsigned n, object **args); } expr mk_bool_true(); @@ -560,7 +560,7 @@ optional reduce_native(environment const & env, expr const & e) { expr const & arg = app_arg(e); if (!is_constant(arg)) return none_expr(); if (app_fn(e) == *g_lean_reduce_bool) { - object * r = ir::run_boxed(env, options(), const_name(arg), 0, nullptr); + object * r = ir::run_boxed_kernel(env, options(), const_name(arg), 0, nullptr); if (!lean_is_scalar(r)) { lean_dec_ref(r); throw kernel_exception(env, "type checker failure, unexpected result value for 'Lean.reduceBool'"); @@ -568,7 +568,7 @@ optional reduce_native(environment const & env, expr const & e) { return lean_unbox(r) == 0 ? some_expr(mk_bool_false()) : some_expr(mk_bool_true()); } if (app_fn(e) == *g_lean_reduce_nat) { - object * r = ir::run_boxed(env, options(), const_name(arg), 0, nullptr); + object * r = ir::run_boxed_kernel(env, options(), const_name(arg), 0, nullptr); if (lean_is_scalar(r) || lean_is_mpz(r)) { return some_expr(mk_lit(literal(nat(r)))); } else { @@ -1193,24 +1193,6 @@ type_checker::~type_checker() { delete m_st; } -extern "C" LEAN_EXPORT lean_object * lean_kernel_is_def_eq(lean_object * env, lean_object * lctx, lean_object * a, lean_object * b) { - return catch_kernel_exceptions([&]() { - return lean_box(type_checker(environment(env), local_ctx(lctx)).is_def_eq(expr(a), expr(b))); - }); -} - -extern "C" LEAN_EXPORT lean_object * lean_kernel_whnf(lean_object * env, lean_object * lctx, lean_object * a) { - return catch_kernel_exceptions([&]() { - return type_checker(environment(env), local_ctx(lctx)).whnf(expr(a)).steal(); - }); -} - -extern "C" LEAN_EXPORT lean_object * lean_kernel_check(lean_object * env, lean_object * lctx, lean_object * a) { - return catch_kernel_exceptions([&]() { - return type_checker(environment(env), local_ctx(lctx)).check(expr(a)).steal(); - }); -} - inline static expr * new_persistent_expr_const(name const & n) { expr * e = new expr(mk_const(n)); mark_persistent(e->raw()); diff --git a/stage0/src/kernel/type_checker.h b/stage0/src/kernel/type_checker.h index 3ab2ce643b72..ac387564aa75 100644 --- a/stage0/src/kernel/type_checker.h +++ b/stage0/src/kernel/type_checker.h @@ -9,6 +9,7 @@ Author: Leonardo de Moura #include #include #include +#include "runtime/flet.h" #include "util/lbool.h" #include "util/name_set.h" #include "util/name_generator.h" @@ -104,6 +105,7 @@ class type_checker { optional reduce_pow(expr const & e); optional reduce_nat(expr const & e); public: + // The following two constructor are used only by the old compiler and should be deleted with it type_checker(state & st, local_ctx const & lctx, definition_safety ds = definition_safety::safe); type_checker(state & st, definition_safety ds = definition_safety::safe):type_checker(st, local_ctx(), ds) {} type_checker(environment const & env, local_ctx const & lctx, diagnostics * diag = nullptr, definition_safety ds = definition_safety::safe); diff --git a/stage0/src/library/CMakeLists.txt b/stage0/src/library/CMakeLists.txt index 7fbbb41851eb..30c2c8924af4 100644 --- a/stage0/src/library/CMakeLists.txt +++ b/stage0/src/library/CMakeLists.txt @@ -6,4 +6,5 @@ add_library(library OBJECT expr_lt.cpp projection.cpp aux_recursors.cpp profiling.cpp time_task.cpp - formatter.cpp) + formatter.cpp + elab_environment.cpp) diff --git a/stage0/src/library/aux_recursors.cpp b/stage0/src/library/aux_recursors.cpp index 724e8c2b1044..d4134090fb75 100644 --- a/stage0/src/library/aux_recursors.cpp +++ b/stage0/src/library/aux_recursors.cpp @@ -10,11 +10,11 @@ namespace lean { extern "C" uint8 lean_is_aux_recursor(object * env, object * n); extern "C" uint8 lean_is_no_confusion(object * env, object * n); -bool is_aux_recursor(environment const & env, name const & r) { +bool is_aux_recursor(elab_environment const & env, name const & r) { return lean_is_aux_recursor(env.to_obj_arg(), r.to_obj_arg()); } -bool is_no_confusion(environment const & env, name const & r) { +bool is_no_confusion(elab_environment const & env, name const & r) { return lean_is_no_confusion(env.to_obj_arg(), r.to_obj_arg()); } } diff --git a/stage0/src/library/aux_recursors.h b/stage0/src/library/aux_recursors.h index 49556be1f862..1d66ee423ff7 100644 --- a/stage0/src/library/aux_recursors.h +++ b/stage0/src/library/aux_recursors.h @@ -5,11 +5,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { /** \brief Return true iff \c n is the name of an auxiliary recursor. \see add_aux_recursor */ -bool is_aux_recursor(environment const & env, name const & n); -bool is_no_confusion(environment const & env, name const & n); +bool is_aux_recursor(elab_environment const & env, name const & n); +bool is_no_confusion(elab_environment const & env, name const & n); } diff --git a/stage0/src/library/class.cpp b/stage0/src/library/class.cpp index f67e89f2cb59..e37bdaf7b3e2 100644 --- a/stage0/src/library/class.cpp +++ b/stage0/src/library/class.cpp @@ -15,9 +15,9 @@ extern "C" uint8 lean_is_out_param(object* e); extern "C" uint8 lean_has_out_params(object* env, object* n); bool is_class_out_param(expr const & e) { return lean_is_out_param(e.to_obj_arg()); } -bool has_class_out_params(environment const & env, name const & c) { return lean_has_out_params(env.to_obj_arg(), c.to_obj_arg()); } -bool is_class(environment const & env, name const & c) { return lean_is_class(env.to_obj_arg(), c.to_obj_arg()); } -bool is_instance(environment const & env, name const & i) { return lean_is_instance(env.to_obj_arg(), i.to_obj_arg()); } +bool has_class_out_params(elab_environment const & env, name const & c) { return lean_has_out_params(env.to_obj_arg(), c.to_obj_arg()); } +bool is_class(elab_environment const & env, name const & c) { return lean_is_class(env.to_obj_arg(), c.to_obj_arg()); } +bool is_instance(elab_environment const & env, name const & i) { return lean_is_instance(env.to_obj_arg(), i.to_obj_arg()); } static name * g_anonymous_inst_name_prefix = nullptr; diff --git a/stage0/src/library/class.h b/stage0/src/library/class.h index ed8768cfc845..2dd2c21c7d13 100644 --- a/stage0/src/library/class.h +++ b/stage0/src/library/class.h @@ -6,11 +6,12 @@ Author: Leonardo de Moura */ #pragma once #include "library/util.h" +#include "library/elab_environment.h" namespace lean { /** \brief Return true iff \c c was declared with \c add_class. */ -bool is_class(environment const & env, name const & c); +bool is_class(elab_environment const & env, name const & c); /** \brief Return true iff \c i was declared with \c add_instance. */ -bool is_instance(environment const & env, name const & i); +bool is_instance(elab_environment const & env, name const & i); name const & get_anonymous_instance_prefix(); name mk_anonymous_inst_name(unsigned idx); @@ -20,7 +21,7 @@ bool is_anonymous_inst_name(name const & n); bool is_class_out_param(expr const & e); /** \brief Return true iff c is a type class that contains an `outParam` */ -bool has_class_out_params(environment const & env, name const & c); +bool has_class_out_params(elab_environment const & env, name const & c); void initialize_class(); void finalize_class(); diff --git a/stage0/src/library/compiler/closed_term_cache.cpp b/stage0/src/library/compiler/closed_term_cache.cpp index 2dbdfa40dd59..8995a95d84e9 100644 --- a/stage0/src/library/compiler/closed_term_cache.cpp +++ b/stage0/src/library/compiler/closed_term_cache.cpp @@ -5,16 +5,17 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #include "library/util.h" +#include "library/elab_environment.h" namespace lean { extern "C" object * lean_cache_closed_term_name(object * env, object * e, object * n); extern "C" object * lean_get_closed_term_name(object * env, object * e); -optional get_closed_term_name(environment const & env, expr const & e) { +optional get_closed_term_name(elab_environment const & env, expr const & e) { return to_optional(lean_get_closed_term_name(env.to_obj_arg(), e.to_obj_arg())); } -environment cache_closed_term_name(environment const & env, expr const & e, name const & n) { - return environment(lean_cache_closed_term_name(env.to_obj_arg(), e.to_obj_arg(), n.to_obj_arg())); +elab_environment cache_closed_term_name(elab_environment const & env, expr const & e, name const & n) { + return elab_environment(lean_cache_closed_term_name(env.to_obj_arg(), e.to_obj_arg(), n.to_obj_arg())); } } diff --git a/stage0/src/library/compiler/closed_term_cache.h b/stage0/src/library/compiler/closed_term_cache.h index a1397a9c6449..e4b2b506b8a1 100644 --- a/stage0/src/library/compiler/closed_term_cache.h +++ b/stage0/src/library/compiler/closed_term_cache.h @@ -5,8 +5,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -optional get_closed_term_name(environment const & env, expr const & e); -environment cache_closed_term_name(environment const & env, expr const & e, name const & n); +optional get_closed_term_name(elab_environment const & env, expr const & e); +elab_environment cache_closed_term_name(elab_environment const & env, expr const & e, name const & n); } diff --git a/stage0/src/library/compiler/compiler.cpp b/stage0/src/library/compiler/compiler.cpp index eb75c93a2737..bed22bb5a599 100644 --- a/stage0/src/library/compiler/compiler.cpp +++ b/stage0/src/library/compiler/compiler.cpp @@ -44,19 +44,19 @@ static name get_real_name(name const & n) { return n; } -static comp_decls to_comp_decls(environment const & env, names const & cs) { +static comp_decls to_comp_decls(elab_environment const & env, names const & cs) { bool allow_opaque = true; return map2(cs, [&](name const & n) { return comp_decl(get_real_name(n), env.get(n).get_value(allow_opaque)); }); } -static expr eta_expand(environment const & env, expr const & e) { - return type_checker(env).eta_expand(e); +static expr eta_expand(elab_environment const & env, expr const & e) { + return type_checker(env.to_kernel_env()).eta_expand(e); } template -comp_decls apply(F && f, environment const & env, comp_decls const & ds) { +comp_decls apply(F && f, elab_environment const & env, comp_decls const & ds) { return map(ds, [&](comp_decl const & d) { return comp_decl(d.fst(), f(env, d.snd())); }); } @@ -75,7 +75,7 @@ void trace_comp_decls(comp_decls const & ds) { } } -static environment cache_stage1(environment env, comp_decls const & ds) { +static elab_environment cache_stage1(elab_environment env, comp_decls const & ds) { for (comp_decl const & d : ds) { name n = d.fst(); expr v = d.snd(); @@ -94,7 +94,7 @@ static expr ensure_arity(expr const & t, unsigned arity) { return update_binding(t, binding_domain(t), ensure_arity(binding_body(t), arity-1)); } -static environment cache_stage2(environment env, comp_decls const & ds, bool only_new_ones = false) { +static elab_environment cache_stage2(elab_environment env, comp_decls const & ds, bool only_new_ones = false) { buffer ts; ll_infer_type(env, ds, ts); lean_assert(ts.size() == length(ds)); @@ -116,11 +116,11 @@ static environment cache_stage2(environment env, comp_decls const & ds, bool onl } /* Cache the declarations in `ds` that have not already been cached. */ -static environment cache_new_stage2(environment env, comp_decls const & ds) { +static elab_environment cache_new_stage2(elab_environment env, comp_decls const & ds) { return cache_stage2(env, ds, true); } -bool is_main_fn(environment const & env, name const & n) { +bool is_main_fn(elab_environment const & env, name const & n) { if (n == "main") return true; if (optional c = get_export_name_for(env, n)) { return *c == "main"; @@ -160,15 +160,15 @@ bool is_main_fn_type(expr const & type) { extern "C" object* lean_csimp_replace_constants(object* env, object* n); -expr csimp_replace_constants(environment const & env, expr const & e) { +expr csimp_replace_constants(elab_environment const & env, expr const & e) { return expr(lean_csimp_replace_constants(env.to_obj_arg(), e.to_obj_arg())); } -bool is_matcher(environment const & env, comp_decls const & ds) { +bool is_matcher(elab_environment const & env, comp_decls const & ds) { return length(ds) == 1 && is_matcher(env, head(ds).fst()); } -environment compile(environment const & env, options const & opts, names cs) { +elab_environment compile(elab_environment const & env, options const & opts, names cs) { /* Do not generate code for irrelevant decls */ cs = filter(cs, [&](name const & c) { return !is_irrelevant_type(env, env.get(c).get_type());}); if (empty(cs)) return env; @@ -202,8 +202,8 @@ environment compile(environment const & env, options const & opts, names cs) { csimp_cfg cfg(opts); // Use the following line to see compiler intermediate steps // scope_traces_as_string trace_scope; - auto simp = [&](environment const & env, expr const & e) { return csimp(env, e, cfg); }; - auto esimp = [&](environment const & env, expr const & e) { return cesimp(env, e, cfg); }; + auto simp = [&](elab_environment const & env, expr const & e) { return csimp(env, e, cfg); }; + auto esimp = [&](elab_environment const & env, expr const & e) { return cesimp(env, e, cfg); }; trace_compiler(name({"compiler", "input"}), ds); ds = apply(eta_expand, env, ds); trace_compiler(name({"compiler", "eta_expand"}), ds); @@ -218,7 +218,7 @@ environment compile(environment const & env, options const & opts, names cs) { ds = apply(simp, env, ds); trace_compiler(name({"compiler", "simp"}), ds); // trace(ds); - environment new_env = env; + elab_environment new_env = env; std::tie(new_env, ds) = eager_lambda_lifting(new_env, ds, cfg); trace_compiler(name({"compiler", "eager_lambda_lifting"}), ds); ds = apply(max_sharing, ds); @@ -274,8 +274,8 @@ environment compile(environment const & env, options const & opts, names cs) { } extern "C" LEAN_EXPORT object * lean_compile_decls(object * env, object * opts, object * decls) { - return catch_kernel_exceptions([&]() { - return compile(environment(env), options(opts, true), names(decls, true)); + return catch_kernel_exceptions([&]() { + return compile(elab_environment(env), options(opts, true), names(decls, true)); }); } diff --git a/stage0/src/library/compiler/compiler.h b/stage0/src/library/compiler/compiler.h index 16c1fe8b3b6e..d620c53aa644 100644 --- a/stage0/src/library/compiler/compiler.h +++ b/stage0/src/library/compiler/compiler.h @@ -5,10 +5,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -environment compile(environment const & env, options const & opts, names cs); -inline environment compile(environment const & env, options const & opts, name const & c) { +elab_environment compile(elab_environment const & env, options const & opts, names cs); +inline elab_environment compile(elab_environment const & env, options const & opts, name const & c) { return compile(env, opts, names(c)); } void initialize_compiler(); diff --git a/stage0/src/library/compiler/cse.cpp b/stage0/src/library/compiler/cse.cpp index 057b30356796..44db877b0c8c 100644 --- a/stage0/src/library/compiler/cse.cpp +++ b/stage0/src/library/compiler/cse.cpp @@ -8,7 +8,7 @@ Author: Leonardo de Moura #include #include "runtime/flet.h" #include "util/name_generator.h" -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "kernel/instantiate.h" #include "kernel/abstract.h" #include "kernel/for_each_fn.h" @@ -21,7 +21,7 @@ namespace lean { static name * g_cse_fresh = nullptr; class cse_fn { - environment m_env; + elab_environment m_env; name_generator m_ngen; bool m_before_erasure; expr_map m_map; @@ -148,14 +148,14 @@ class cse_fn { } public: - cse_fn(environment const & env, bool before_erasure): + cse_fn(elab_environment const & env, bool before_erasure): m_env(env), m_ngen(*g_cse_fresh), m_before_erasure(before_erasure) { } expr operator()(expr const & e) { return visit(e); } }; -expr cse_core(environment const & env, expr const & e, bool before_erasure) { +expr cse_core(elab_environment const & env, expr const & e, bool before_erasure) { return cse_fn(env, before_erasure)(e); } @@ -170,6 +170,7 @@ expr cse_core(environment const & env, expr const & e, bool before_erasure) { ``` The "else"-branch is duplicated by the equation compiler for each constructor different from `expr.app`. */ class cce_fn { + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; buffer m_fvars; @@ -178,7 +179,7 @@ class cce_fn { name m_j; unsigned m_next_idx{1}; public: - environment & env() { return m_st.env(); } + elab_environment & env() { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -394,8 +395,8 @@ class cce_fn { } public: - cce_fn(environment const & env, local_ctx const & lctx): - m_st(env), m_lctx(lctx), m_j("_j") { + cce_fn(elab_environment const & env, local_ctx const & lctx): + m_env(env), m_st(env), m_lctx(lctx), m_j("_j") { } expr operator()(expr const & e) { @@ -404,7 +405,7 @@ class cce_fn { } }; -expr cce_core(environment const & env, local_ctx const & lctx, expr const & e) { +expr cce_core(elab_environment const & env, local_ctx const & lctx, expr const & e) { return cce_fn(env, lctx)(e); } diff --git a/stage0/src/library/compiler/cse.h b/stage0/src/library/compiler/cse.h index 82bd830a7588..4488f16538fe 100644 --- a/stage0/src/library/compiler/cse.h +++ b/stage0/src/library/compiler/cse.h @@ -5,15 +5,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { /* Common subexpression elimination */ -expr cse_core(environment const & env, expr const & e, bool before_erasure); -inline expr cse(environment const & env, expr const & e) { return cse_core(env, e, true); } -inline expr ecse(environment const & env, expr const & e) { return cse_core(env, e, false); } +expr cse_core(elab_environment const & env, expr const & e, bool before_erasure); +inline expr cse(elab_environment const & env, expr const & e) { return cse_core(env, e, true); } +inline expr ecse(elab_environment const & env, expr const & e) { return cse_core(env, e, false); } /* Common case elimination */ -expr cce_core(environment const & env, local_ctx const & lctx, expr const & e); -inline expr cce(environment const & env, expr const & e) { return cce_core(env, local_ctx(), e); } +expr cce_core(elab_environment const & env, local_ctx const & lctx, expr const & e); +inline expr cce(elab_environment const & env, expr const & e) { return cce_core(env, local_ctx(), e); } void initialize_cse(); void finalize_cse(); } diff --git a/stage0/src/library/compiler/csimp.cpp b/stage0/src/library/compiler/csimp.cpp index 35bb6c375446..1bc325dac84d 100644 --- a/stage0/src/library/compiler/csimp.cpp +++ b/stage0/src/library/compiler/csimp.cpp @@ -64,6 +64,7 @@ optional fold_bin_op(bool before_erasure, expr const & f, expr const & a, class csimp_fn { typedef expr_pair_struct_map jp_cache; + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; bool m_before_erasure; @@ -96,7 +97,7 @@ class csimp_fn { typedef rb_expr_map expr2ctor; expr2ctor m_expr2ctor; - environment const & env() const { return m_st.env(); } + elab_environment const & env() const { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -1625,12 +1626,12 @@ class csimp_fn { } struct is_recursive_fn { - environment const & m_env; + elab_environment const & m_env; csimp_cfg const & m_cfg; bool m_before_erasure; name m_target; - is_recursive_fn(environment const & env, csimp_cfg const & cfg, bool before_erasure): + is_recursive_fn(elab_environment const & env, csimp_cfg const & cfg, bool before_erasure): m_env(env), m_cfg(cfg), m_before_erasure(before_erasure) { } @@ -1969,8 +1970,8 @@ class csimp_fn { } public: - csimp_fn(environment const & env, local_ctx const & lctx, bool before_erasure, csimp_cfg const & cfg): - m_st(env), m_lctx(lctx), m_before_erasure(before_erasure), m_cfg(cfg), m_x("_x"), m_j("j") {} + csimp_fn(elab_environment const & env, local_ctx const & lctx, bool before_erasure, csimp_cfg const & cfg): + m_env(env), m_st(env), m_lctx(lctx), m_before_erasure(before_erasure), m_cfg(cfg), m_x("_x"), m_j("j") {} expr operator()(expr const & e) { if (is_lambda(e)) { @@ -1992,7 +1993,7 @@ bool at_most_once(expr const & e, name const & x) { /* Eliminate join-points that are used only once */ class elim_jp1_fn { - environment const & m_env; + elab_environment const & m_env; local_ctx m_lctx; bool m_before_erasure; name_generator m_ngen; @@ -2093,7 +2094,7 @@ class elim_jp1_fn { } public: - elim_jp1_fn(environment const & env, local_ctx const & lctx, bool before_erasure): + elim_jp1_fn(elab_environment const & env, local_ctx const & lctx, bool before_erasure): m_env(env), m_lctx(lctx), m_before_erasure(before_erasure) {} expr operator()(expr const & e) { m_expanded = false; @@ -2103,7 +2104,7 @@ class elim_jp1_fn { bool expanded() const { return m_expanded; } }; -expr csimp_core(environment const & env, local_ctx const & lctx, expr const & e0, bool before_erasure, csimp_cfg const & cfg) { +expr csimp_core(elab_environment const & env, local_ctx const & lctx, expr const & e0, bool before_erasure, csimp_cfg const & cfg) { csimp_fn simp(env, lctx, before_erasure, cfg); elim_jp1_fn elim_jp1(env, lctx, before_erasure); expr e = e0; diff --git a/stage0/src/library/compiler/csimp.h b/stage0/src/library/compiler/csimp.h index db3746a770d4..a4ef44b7bd0c 100644 --- a/stage0/src/library/compiler/csimp.h +++ b/stage0/src/library/compiler/csimp.h @@ -5,7 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { struct csimp_cfg { /* If `m_inline` == false, then we will not inline `c` even if it is marked with the attribute `[inline]`. */ @@ -23,11 +23,11 @@ struct csimp_cfg { csimp_cfg(); }; -expr csimp_core(environment const & env, local_ctx const & lctx, expr const & e, bool before_erasure, csimp_cfg const & cfg); -inline expr csimp(environment const & env, expr const & e, csimp_cfg const & cfg = csimp_cfg()) { +expr csimp_core(elab_environment const & env, local_ctx const & lctx, expr const & e, bool before_erasure, csimp_cfg const & cfg); +inline expr csimp(elab_environment const & env, expr const & e, csimp_cfg const & cfg = csimp_cfg()) { return csimp_core(env, local_ctx(), e, true, cfg); } -inline expr cesimp(environment const & env, expr const & e, csimp_cfg const & cfg = csimp_cfg()) { +inline expr cesimp(elab_environment const & env, expr const & e, csimp_cfg const & cfg = csimp_cfg()) { return csimp_core(env, local_ctx(), e, false, cfg); } } diff --git a/stage0/src/library/compiler/eager_lambda_lifting.cpp b/stage0/src/library/compiler/eager_lambda_lifting.cpp index 2f838dad1f2a..3e7d1da698dc 100644 --- a/stage0/src/library/compiler/eager_lambda_lifting.cpp +++ b/stage0/src/library/compiler/eager_lambda_lifting.cpp @@ -98,6 +98,7 @@ static bool depends_on_fvar(local_ctx const & lctx, buffer const & params, Remark: we also skip this transformation for definitions marked as `[inline]` or `[instance]`. */ class eager_lambda_lifting_fn { + elab_environment m_env; type_checker::state m_st; csimp_cfg m_cfg; local_ctx m_lctx; @@ -108,7 +109,7 @@ class eager_lambda_lifting_fn { name_set m_nonterminal_lambdas; unsigned m_next_idx{1}; - environment const & env() const { return m_st.env(); } + elab_environment const & env() const { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -234,12 +235,13 @@ class eager_lambda_lifting_fn { } expr type = cheap_beta_reduce(type_checker(m_st).infer(code)); name n = next_name(); - /* We add the auxiliary declaration `n` as a "meta" axiom to the environment. + /* We add the auxiliary declaration `n` as a "meta" axiom to the elab_environment. This is a hack to make sure we can use `csimp` to simplify `code` and other definitions that use `n`. We used a similar hack at `specialize.cpp`. */ declaration aux_ax = mk_axiom(n, names(), type, true /* meta */); - m_st.env() = env().add(aux_ax, false); + m_env = m_env.add(aux_ax, false); + m_st.env() = m_env; m_new_decls.push_back(comp_decl(n, code)); return mk_app(mk_constant(n), new_params); } catch (exception &) { @@ -398,10 +400,10 @@ class eager_lambda_lifting_fn { } public: - eager_lambda_lifting_fn(environment const & env, csimp_cfg const & cfg): - m_st(env), m_cfg(cfg) {} + eager_lambda_lifting_fn(elab_environment const & env, csimp_cfg const & cfg): + m_env(env), m_st(env), m_cfg(cfg) {} - pair operator()(comp_decl const & cdecl) { + pair operator()(comp_decl const & cdecl) { m_base_name = cdecl.fst(); expr r = visit(cdecl.snd(), true); comp_decl new_cdecl(cdecl.fst(), r); @@ -410,7 +412,7 @@ class eager_lambda_lifting_fn { } }; -pair eager_lambda_lifting(environment env, comp_decls const & ds, csimp_cfg const & cfg) { +pair eager_lambda_lifting(elab_environment env, comp_decls const & ds, csimp_cfg const & cfg) { comp_decls r; for (comp_decl const & d : ds) { if (has_inline_attribute(env, d.fst()) || is_instance(env, d.fst())) { diff --git a/stage0/src/library/compiler/eager_lambda_lifting.h b/stage0/src/library/compiler/eager_lambda_lifting.h index 04ab99872daa..bccc9ba42b0f 100644 --- a/stage0/src/library/compiler/eager_lambda_lifting.h +++ b/stage0/src/library/compiler/eager_lambda_lifting.h @@ -5,11 +5,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "library/compiler/csimp.h" namespace lean { /** \brief Eager version of lambda lifting. See comment at eager_lambda_lifting.cpp. */ -pair eager_lambda_lifting(environment env, comp_decls const & ds, csimp_cfg const & cfg); +pair eager_lambda_lifting(elab_environment env, comp_decls const & ds, csimp_cfg const & cfg); /* Return true iff `fn` is the name of an auxiliary function generated by `eager_lambda_lifting`. */ bool is_elambda_lifting_name(name fn); }; diff --git a/stage0/src/library/compiler/erase_irrelevant.cpp b/stage0/src/library/compiler/erase_irrelevant.cpp index 42057654f34a..727ec117cc5c 100644 --- a/stage0/src/library/compiler/erase_irrelevant.cpp +++ b/stage0/src/library/compiler/erase_irrelevant.cpp @@ -16,6 +16,7 @@ Author: Leonardo de Moura namespace lean { class erase_irrelevant_fn { typedef std::tuple let_entry; + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; buffer m_let_fvars; @@ -24,7 +25,7 @@ class erase_irrelevant_fn { unsigned m_next_idx{1}; expr_map m_irrelevant_cache; - environment & env() { return m_st.env(); } + elab_environment & env() { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -485,14 +486,14 @@ class erase_irrelevant_fn { lean_unreachable(); } public: - erase_irrelevant_fn(environment const & env, local_ctx const & lctx): - m_st(env), m_lctx(lctx), m_x("_x") {} + erase_irrelevant_fn(elab_environment const & env, local_ctx const & lctx): + m_env(env), m_st(env), m_lctx(lctx), m_x("_x") {} expr operator()(expr const & e) { return mk_let(0, visit(e)); } }; -expr erase_irrelevant_core(environment const & env, local_ctx const & lctx, expr const & e) { +expr erase_irrelevant_core(elab_environment const & env, local_ctx const & lctx, expr const & e) { return erase_irrelevant_fn(env, lctx)(e); } } diff --git a/stage0/src/library/compiler/erase_irrelevant.h b/stage0/src/library/compiler/erase_irrelevant.h index 5ced49263321..15f3ffd7796d 100644 --- a/stage0/src/library/compiler/erase_irrelevant.h +++ b/stage0/src/library/compiler/erase_irrelevant.h @@ -5,8 +5,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -expr erase_irrelevant_core(environment const & env, local_ctx const & lctx, expr const & e); -inline expr erase_irrelevant(environment const & env, expr const & e) { return erase_irrelevant_core(env, local_ctx(), e); } +expr erase_irrelevant_core(elab_environment const & env, local_ctx const & lctx, expr const & e); +inline expr erase_irrelevant(elab_environment const & env, expr const & e) { return erase_irrelevant_core(env, local_ctx(), e); } } diff --git a/stage0/src/library/compiler/export_attribute.cpp b/stage0/src/library/compiler/export_attribute.cpp index 5ab58f6cdfc6..6388a77d97fd 100644 --- a/stage0/src/library/compiler/export_attribute.cpp +++ b/stage0/src/library/compiler/export_attribute.cpp @@ -6,10 +6,11 @@ Author: Leonardo de Moura */ #include "library/constants.h" #include "library/util.h" +#include "library/elab_environment.h" namespace lean { extern "C" object* lean_get_export_name_for(object* env, object *n); -optional get_export_name_for(environment const & env, name const & n) { +optional get_export_name_for(elab_environment const & env, name const & n) { return to_optional(lean_get_export_name_for(env.to_obj_arg(), n.to_obj_arg())); } } diff --git a/stage0/src/library/compiler/export_attribute.h b/stage0/src/library/compiler/export_attribute.h index 81744fdd800a..06026945c017 100644 --- a/stage0/src/library/compiler/export_attribute.h +++ b/stage0/src/library/compiler/export_attribute.h @@ -5,8 +5,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -optional get_export_name_for(environment const & env, name const & n); -inline bool has_export_name(environment const & env, name const & n) { return static_cast(get_export_name_for(env, n)); } +optional get_export_name_for(elab_environment const & env, name const & n); +inline bool has_export_name(elab_environment const & env, name const & n) { return static_cast(get_export_name_for(env, n)); } } diff --git a/stage0/src/library/compiler/extern_attribute.cpp b/stage0/src/library/compiler/extern_attribute.cpp index 447f8929209d..5c00d3a89fe9 100644 --- a/stage0/src/library/compiler/extern_attribute.cpp +++ b/stage0/src/library/compiler/extern_attribute.cpp @@ -23,15 +23,15 @@ Authors: Leonardo de Moura namespace lean { extern "C" object* lean_get_extern_attr_data(object* env, object* n); -optional get_extern_attr_data(environment const & env, name const & fn) { +optional get_extern_attr_data(elab_environment const & env, name const & fn) { return to_optional(lean_get_extern_attr_data(env.to_obj_arg(), fn.to_obj_arg())); } -bool is_extern_constant(environment const & env, name const & c) { +bool is_extern_constant(elab_environment const & env, name const & c) { return static_cast(get_extern_attr_data(env, c)); } -bool is_extern_or_init_constant(environment const & env, name const & c) { +bool is_extern_or_init_constant(elab_environment const & env, name const & c) { if (is_extern_constant(env, c)) { return true; } else if (auto info = env.find(c)) { @@ -44,7 +44,7 @@ bool is_extern_or_init_constant(environment const & env, name const & c) { extern "C" object * lean_get_extern_const_arity(object* env, object*, object* w); -optional get_extern_constant_arity(environment const & env, name const & c) { +optional get_extern_constant_arity(elab_environment const & env, name const & c) { auto arity = get_io_result>(lean_get_extern_const_arity(env.to_obj_arg(), c.to_obj_arg(), lean_io_mk_world())); if (optional aux = arity.get()) { return optional(aux->get_small_value()); @@ -53,7 +53,7 @@ optional get_extern_constant_arity(environment const & env, name const } } -bool get_extern_borrowed_info(environment const & env, name const & c, buffer & borrowed_args, bool & borrowed_res) { +bool get_extern_borrowed_info(elab_environment const & env, name const & c, buffer & borrowed_args, bool & borrowed_res) { if (is_extern_constant(env, c)) { /* Extract borrowed info from type */ expr type = env.get(c).get_type(); @@ -80,7 +80,7 @@ bool get_extern_borrowed_info(environment const & env, name const & c, buffer get_extern_constant_ll_type(environment const & env, name const & c) { +optional get_extern_constant_ll_type(elab_environment const & env, name const & c) { if (is_extern_or_init_constant(env, c)) { unsigned arity = 0; expr type = env.get(c).get_type(); diff --git a/stage0/src/library/compiler/extern_attribute.h b/stage0/src/library/compiler/extern_attribute.h index 84c6d3194bfa..c254121630e2 100644 --- a/stage0/src/library/compiler/extern_attribute.h +++ b/stage0/src/library/compiler/extern_attribute.h @@ -6,15 +6,15 @@ Authors: Leonardo de Moura */ #pragma once #include -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -bool is_extern_constant(environment const & env, name const & c); -bool is_extern_or_init_constant(environment const & env, name const & c); -optional get_extern_constant_ll_type(environment const & env, name const & c); -optional get_extern_constant_arity(environment const & env, name const & c); +bool is_extern_constant(elab_environment const & env, name const & c); +bool is_extern_or_init_constant(elab_environment const & env, name const & c); +optional get_extern_constant_ll_type(elab_environment const & env, name const & c); +optional get_extern_constant_arity(elab_environment const & env, name const & c); typedef object_ref extern_attr_data_value; -optional get_extern_attr_data(environment const & env, name const & c); +optional get_extern_attr_data(elab_environment const & env, name const & c); /* Return true if `c` is an extern constant, and store in borrowed_args and borrowed_res which arguments/results are marked as borrowed. */ -bool get_extern_borrowed_info(environment const & env, name const & c, buffer & borrowed_args, bool & borrowed_res); +bool get_extern_borrowed_info(elab_environment const & env, name const & c, buffer & borrowed_args, bool & borrowed_res); } diff --git a/stage0/src/library/compiler/extract_closed.cpp b/stage0/src/library/compiler/extract_closed.cpp index 929b49e6a3cd..234ce10ecfbd 100644 --- a/stage0/src/library/compiler/extract_closed.cpp +++ b/stage0/src/library/compiler/extract_closed.cpp @@ -25,7 +25,7 @@ bool is_extract_closed_aux_fn(name const & n) { } class extract_closed_fn { - environment m_env; + elab_environment m_env; comp_decls m_input_decls; name_generator m_ngen; local_ctx m_lctx; @@ -34,7 +34,7 @@ class extract_closed_fn { unsigned m_next_idx{1}; expr_map m_closed; - environment const & env() const { return m_env; } + elab_environment const & env() const { return m_env; } name_generator & ngen() { return m_ngen; } name next_name() { @@ -286,11 +286,11 @@ class extract_closed_fn { } public: - extract_closed_fn(environment const & env, comp_decls const & ds): + extract_closed_fn(elab_environment const & env, comp_decls const & ds): m_env(env), m_input_decls(ds) { } - pair operator()(comp_decl const & d) { + pair operator()(comp_decl const & d) { if (arity_was_reduced(d)) { /* Do nothing since `d` will be inlined. */ return mk_pair(env(), comp_decls(d)); @@ -308,11 +308,11 @@ class extract_closed_fn { } }; -pair extract_closed_core(environment const & env, comp_decls const & input_ds, comp_decl const & d) { +pair extract_closed_core(elab_environment const & env, comp_decls const & input_ds, comp_decl const & d) { return extract_closed_fn(env, input_ds)(d); } -pair extract_closed(environment env, comp_decls const & ds) { +pair extract_closed(elab_environment env, comp_decls const & ds) { comp_decls r; for (comp_decl const & d : ds) { comp_decls new_ds; diff --git a/stage0/src/library/compiler/extract_closed.h b/stage0/src/library/compiler/extract_closed.h index 1c8b39d4585f..5c15fdb910b3 100644 --- a/stage0/src/library/compiler/extract_closed.h +++ b/stage0/src/library/compiler/extract_closed.h @@ -5,9 +5,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "library/compiler/util.h" namespace lean { bool is_extract_closed_aux_fn(name const & n); -pair extract_closed(environment env, comp_decls const & ds); +pair extract_closed(elab_environment env, comp_decls const & ds); } diff --git a/stage0/src/library/compiler/find_jp.cpp b/stage0/src/library/compiler/find_jp.cpp index 343128c0c519..a6279e953ef0 100644 --- a/stage0/src/library/compiler/find_jp.cpp +++ b/stage0/src/library/compiler/find_jp.cpp @@ -13,7 +13,7 @@ namespace lean { /* Find join-points */ class find_jp_fn { - environment const & m_env; + elab_environment const & m_env; local_ctx m_lctx; name_generator m_ngen; name_map m_candidates; @@ -130,7 +130,7 @@ class find_jp_fn { } public: - find_jp_fn(environment const & env): + find_jp_fn(elab_environment const & env): m_env(env) {} expr operator()(expr const & e) { @@ -138,7 +138,7 @@ class find_jp_fn { } }; -expr find_jp(environment const & env, expr const & e) { +expr find_jp(elab_environment const & env, expr const & e) { return find_jp_fn(env)(e); } } diff --git a/stage0/src/library/compiler/find_jp.h b/stage0/src/library/compiler/find_jp.h index 2b611a1b59f9..a71e2fcc55c5 100644 --- a/stage0/src/library/compiler/find_jp.h +++ b/stage0/src/library/compiler/find_jp.h @@ -5,7 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -expr find_jp(environment const & env, expr const & e); +expr find_jp(elab_environment const & env, expr const & e); } diff --git a/stage0/src/library/compiler/implemented_by_attribute.cpp b/stage0/src/library/compiler/implemented_by_attribute.cpp index 267c60c4ccb2..8a6afe4ef032 100644 --- a/stage0/src/library/compiler/implemented_by_attribute.cpp +++ b/stage0/src/library/compiler/implemented_by_attribute.cpp @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { extern "C" object* lean_get_implemented_by(object*, object*); -optional get_implemented_by_attribute(environment const & env, name const & n) { +optional get_implemented_by_attribute(elab_environment const & env, name const & n) { return to_optional(lean_get_implemented_by(env.to_obj_arg(), n.to_obj_arg())); } } diff --git a/stage0/src/library/compiler/implemented_by_attribute.h b/stage0/src/library/compiler/implemented_by_attribute.h index 375cb3ba37ba..30b68954d8c6 100644 --- a/stage0/src/library/compiler/implemented_by_attribute.h +++ b/stage0/src/library/compiler/implemented_by_attribute.h @@ -5,11 +5,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -optional get_implemented_by_attribute(environment const & env, name const & n); -inline bool has_implemented_by_attribute(environment const & env, name const & n) { +optional get_implemented_by_attribute(elab_environment const & env, name const & n); +inline bool has_implemented_by_attribute(elab_environment const & env, name const & n) { return static_cast(get_implemented_by_attribute(env, n)); } } diff --git a/stage0/src/library/compiler/init_attribute.cpp b/stage0/src/library/compiler/init_attribute.cpp index eac9ad41f13d..1b5a09986ab8 100644 --- a/stage0/src/library/compiler/init_attribute.cpp +++ b/stage0/src/library/compiler/init_attribute.cpp @@ -5,12 +5,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura */ #include "runtime/object_ref.h" -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { extern "C" object* lean_get_init_fn_name_for(object* env, object* fn); -optional get_init_fn_name_for(environment const & env, name const & n) { +optional get_init_fn_name_for(elab_environment const & env, name const & n) { return to_optional(lean_get_init_fn_name_for(env.to_obj_arg(), n.to_obj_arg())); } } diff --git a/stage0/src/library/compiler/init_attribute.h b/stage0/src/library/compiler/init_attribute.h index 9bfaa62a3d1d..7a0f5707546b 100644 --- a/stage0/src/library/compiler/init_attribute.h +++ b/stage0/src/library/compiler/init_attribute.h @@ -5,11 +5,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -optional get_init_fn_name_for(environment const & env, name const & n); -inline bool has_init_attribute(environment const & env, name const & n) { +optional get_init_fn_name_for(elab_environment const & env, name const & n); +inline bool has_init_attribute(elab_environment const & env, name const & n) { return static_cast(get_init_fn_name_for(env, n)); } } diff --git a/stage0/src/library/compiler/ir.cpp b/stage0/src/library/compiler/ir.cpp index 2c499224f3cd..528b28bf3b42 100644 --- a/stage0/src/library/compiler/ir.cpp +++ b/stage0/src/library/compiler/ir.cpp @@ -100,8 +100,8 @@ std::string decl_to_string(decl const & d) { string_ref r(lean_ir_decl_to_string(d.to_obj_arg())); return r.to_std_string(); } -environment add_decl(environment const & env, decl const & d) { - return environment(lean_ir_add_decl(env.to_obj_arg(), d.to_obj_arg())); +elab_environment add_decl(elab_environment const & env, decl const & d) { + return elab_environment(lean_ir_add_decl(env.to_obj_arg(), d.to_obj_arg())); } } @@ -133,12 +133,13 @@ static ir::type to_ir_type(expr const & e) { } class to_ir_fn { + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; name m_x{"x"}; unsigned m_next_idx{1}; - environment const & env() const { return m_st.env(); } + elab_environment const & env() const { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -493,7 +494,7 @@ class to_ir_fn { return ir::mk_decl(fn, xs, type, b); } public: - to_ir_fn(environment const & env):m_st(env) {} + to_ir_fn(elab_environment const & env):m_env(env), m_st(env) {} ir::decl operator()(comp_decl const & d) { return to_ir_decl(d); } @@ -520,7 +521,7 @@ class to_ir_fn { }; namespace ir { -decl to_ir_decl(environment const & env, comp_decl const & d) { +decl to_ir_decl(elab_environment const & env, comp_decl const & d) { return to_ir_fn(env)(d); } @@ -528,7 +529,7 @@ decl to_ir_decl(environment const & env, comp_decl const & d) { @[export lean.ir.compile_core] def compile (env : Environment) (opts : Options) (decls : Array Decl) : Log × (Except String Environment) := */ -environment compile(environment const & env, options const & opts, comp_decls const & decls) { +elab_environment compile(elab_environment const & env, options const & opts, comp_decls const & decls) { buffer ir_decls; for (comp_decl const & decl : decls) { lean_trace(name({"compiler", "lambda_pure"}), @@ -549,7 +550,7 @@ environment compile(environment const & env, options const & opts, comp_decls co dec_ref(r); throw exception(error.data()); } else { - environment new_env(cnstr_get(v, 0), true); + elab_environment new_env(cnstr_get(v, 0), true); dec_ref(r); return new_env; } @@ -560,28 +561,28 @@ environment compile(environment const & env, options const & opts, comp_decls co def addBoxedVersion (env : Environment) (decl : Decl) : Except String Environment := */ extern "C" object * lean_ir_add_boxed_version(object * env, object * decl); -environment add_boxed_version(environment const & env, decl const & d) { +elab_environment add_boxed_version(elab_environment const & env, decl const & d) { object * v = lean_ir_add_boxed_version(env.to_obj_arg(), d.to_obj_arg()); if (cnstr_tag(v) == 0) { string_ref error(cnstr_get(v, 0), true); dec_ref(v); throw exception(error.data()); } else { - environment new_env(cnstr_get(v, 0), true); + elab_environment new_env(cnstr_get(v, 0), true); dec_ref(v); return new_env; } } -environment add_extern(environment const & env, name const & fn) { +elab_environment add_extern(elab_environment const & env, name const & fn) { decl d = to_ir_fn(env)(fn); - environment new_env = ir::add_decl(env, d); + elab_environment new_env = ir::add_decl(env, d); return add_boxed_version(new_env, d); } extern "C" LEAN_EXPORT object* lean_add_extern(object * env, object * fn) { try { - environment new_env = add_extern(environment(env), name(fn)); + elab_environment new_env = add_extern(elab_environment(env), name(fn)); return mk_except_ok(new_env); } catch (exception & ex) { // throw; // We use to uncomment this line when debugging weird bugs in the Lean/C++ interface. @@ -591,7 +592,7 @@ extern "C" LEAN_EXPORT object* lean_add_extern(object * env, object * fn) { extern "C" object * lean_ir_emit_c(object * env, object * mod_name); -string_ref emit_c(environment const & env, name const & mod_name) { +string_ref emit_c(elab_environment const & env, name const & mod_name) { object * r = lean_ir_emit_c(env.to_obj_arg(), mod_name.to_obj_arg()); string_ref s(cnstr_get(r, 0), true); if (cnstr_tag(r) == 0) { @@ -639,7 +640,7 @@ object_ref to_object_ref(cnstr_info const & info) { } extern "C" LEAN_EXPORT object * lean_ir_get_ctor_layout(object * env0, object * ctor_name0) { - environment const & env = TO_REF(environment, env0); + elab_environment const & env = TO_REF(elab_environment, env0); name const & ctor_name = TO_REF(name, ctor_name0); type_checker::state st(env); try { diff --git a/stage0/src/library/compiler/ir.h b/stage0/src/library/compiler/ir.h index 8a9a092269a2..df91f21ee5b4 100644 --- a/stage0/src/library/compiler/ir.h +++ b/stage0/src/library/compiler/ir.h @@ -6,7 +6,7 @@ Author: Leonardo de Moura */ #pragma once #include -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "library/compiler/util.h" namespace lean { namespace ir { @@ -35,10 +35,10 @@ typedef object_ref decl; typedef object_ref decl; std::string decl_to_string(decl const & d); void test(decl const & d); -environment compile(environment const & env, options const & opts, comp_decls const & decls); -environment add_extern(environment const & env, name const & fn); -LEAN_EXPORT string_ref emit_c(environment const & env, name const & mod_name); -void emit_llvm(environment const & env, name const & mod_name, std::string const &filepath); +elab_environment compile(elab_environment const & env, options const & opts, comp_decls const & decls); +elab_environment add_extern(elab_environment const & env, name const & fn); +LEAN_EXPORT string_ref emit_c(elab_environment const & env, name const & mod_name); +void emit_llvm(elab_environment const & env, name const & mod_name, std::string const &filepath); } void initialize_ir(); void finalize_ir(); diff --git a/stage0/src/library/compiler/ir_interpreter.cpp b/stage0/src/library/compiler/ir_interpreter.cpp index 978aa853f458..2f32ef862150 100644 --- a/stage0/src/library/compiler/ir_interpreter.cpp +++ b/stage0/src/library/compiler/ir_interpreter.cpp @@ -19,7 +19,7 @@ Implementation The interpreter mainly consists of a homogeneous stack of `value`s, which are either unboxed values or pointers to boxed objects. The IR type system tells us which union member is active at any time. IR variables are mapped to stack slots by adding the current base pointer to the variable index. Further stacks are used for storing join points and call -stack metadata. The interpreted IR is taken directly from the environment. Whenever possible, we try to switch to native +stack metadata. The interpreted IR is taken directly from the elab_environment. Whenever possible, we try to switch to native code by checking for the mangled symbol via dlsym/GetProcAddress, which is also how we can call external functions (which only works if the file declaring them has already been compiled). We always call the "boxed" versions of native functions, which have a (relatively) homogeneous ABI that we can use without runtime code generation; see also @@ -182,7 +182,7 @@ type decl_type(decl const & b) { return cnstr_get_type(b, 2); } fn_body const & decl_fun_body(decl const & b) { lean_assert(decl_tag(b) == decl_kind::Fun); return cnstr_get_ref_t(b, 3); } extern "C" object * lean_ir_find_env_decl(object * env, object * n); -option_ref find_ir_decl(environment const & env, name const & n) { +option_ref find_ir_decl(elab_environment const & env, name const & n) { return option_ref(lean_ir_find_env_decl(env.to_obj_arg(), n.to_obj_arg())); } @@ -213,12 +213,12 @@ static bool type_is_scalar(type t) { } extern "C" object* lean_get_regular_init_fn_name_for(object* env, object* fn); -optional get_regular_init_fn_name_for(environment const & env, name const & n) { +optional get_regular_init_fn_name_for(elab_environment const & env, name const & n) { return to_optional(lean_get_regular_init_fn_name_for(env.to_obj_arg(), n.to_obj_arg())); } extern "C" object* lean_get_builtin_init_fn_name_for(object* env, object* fn); -optional get_builtin_init_fn_name_for(environment const & env, name const & n) { +optional get_builtin_init_fn_name_for(elab_environment const & env, name const & n) { return to_optional(lean_get_builtin_init_fn_name_for(env.to_obj_arg(), n.to_obj_arg())); } @@ -355,7 +355,7 @@ class interpreter { frame(name const & mFn, size_t mArgBp, size_t mJpBp) : m_fn(mFn), m_arg_bp(mArgBp), m_jp_bp(mJpBp) {} }; std::vector m_call_stack; - environment const & m_env; + elab_environment const & m_env; options const & m_opts; // if `false`, use IR code where possible bool m_prefer_native; @@ -393,7 +393,7 @@ class interpreter { public: template - static inline T with_interpreter(environment const & env, options const & opts, name const & fn, std::function const & f) { + static inline T with_interpreter(elab_environment const & env, options const & opts, name const & fn, std::function const & f) { if (g_interpreter && is_eqp(g_interpreter->m_env, env) && is_eqp(g_interpreter->m_opts, opts)) { return f(*g_interpreter); } else { @@ -808,7 +808,7 @@ class interpreter { } } - /** \brief Retrieve Lean declaration from environment. */ + /** \brief Retrieve Lean declaration from elab_environment. */ decl get_decl(name const & fn) { option_ref d = find_ir_decl(m_env, fn); if (!d) { @@ -931,7 +931,7 @@ class interpreter { // static closure stub static object * stub_m_aux(object ** args) { - environment env(args[0]); + elab_environment env(args[0]); options opts(args[1]); return with_interpreter(env, opts, decl_fun_id(TO_REF(decl, args[2])), [&](interpreter & interp) { return interp.stub_m(args); @@ -979,7 +979,7 @@ class interpreter { } } public: - explicit interpreter(environment const & env, options const & opts) : m_env(env), m_opts(opts) { + explicit interpreter(elab_environment const & env, options const & opts) : m_env(env), m_opts(opts) { m_prefer_native = opts.get_bool(*g_interpreter_prefer_native, LEAN_DEFAULT_INTERPRETER_PREFER_NATIVE); } @@ -1090,24 +1090,34 @@ class interpreter { extern "C" object * lean_decl_get_sorry_dep(object * env, object * n); -optional get_sorry_dep(environment const & env, name const & n) { +optional get_sorry_dep(elab_environment const & env, name const & n) { return option_ref(lean_decl_get_sorry_dep(env.to_obj_arg(), n.to_obj_arg())).get(); } -object * run_boxed(environment const & env, options const & opts, name const & fn, unsigned n, object **args) { +object * run_boxed(elab_environment const & env, options const & opts, name const & fn, unsigned n, object **args) { if (optional decl_with_sorry = get_sorry_dep(env, fn)) { throw exception(sstream() << "cannot evaluate code because '" << *decl_with_sorry << "' uses 'sorry' and/or contains errors"); } return interpreter::with_interpreter(env, opts, fn, [&](interpreter & interp) { return interp.call_boxed(fn, n, args); }); } -uint32 run_main(environment const & env, options const & opts, int argv, char * argc[]) { + +extern "C" obj_res lean_elab_environment_of_kernel_env(obj_arg); +elab_environment elab_environment_of_kernel_env(environment const & env) { + return elab_environment(lean_elab_environment_of_kernel_env(env.to_obj_arg())); +} + +object * run_boxed_kernel(environment const & env, options const & opts, name const & fn, unsigned n, object **args) { + return run_boxed(elab_environment_of_kernel_env(env), opts, fn, n, args); +} + +uint32 run_main(elab_environment const & env, options const & opts, int argv, char * argc[]) { return interpreter::with_interpreter(env, opts, "main", [&](interpreter & interp) { return interp.run_main(argv, argc); }); } extern "C" LEAN_EXPORT object * lean_eval_const(object * env, object * opts, object * c) { try { - return mk_cnstr(1, run_boxed(TO_REF(environment, env), TO_REF(options, opts), TO_REF(name, c), 0, 0)).steal(); + return mk_cnstr(1, run_boxed(TO_REF(elab_environment, env), TO_REF(options, opts), TO_REF(name, c), 0, 0)).steal(); } catch (exception & ex) { return mk_cnstr(0, string_ref(ex.what())).steal(); } @@ -1134,7 +1144,7 @@ extern "C" LEAN_EXPORT object * lean_run_mod_init(object * mod, object *) { } extern "C" LEAN_EXPORT object * lean_run_init(object * env, object * opts, object * decl, object * init_decl, object *) { - return interpreter::with_interpreter(TO_REF(environment, env), TO_REF(options, opts), TO_REF(name, decl), [&](interpreter & interp) { + return interpreter::with_interpreter(TO_REF(elab_environment, env), TO_REF(options, opts), TO_REF(name, decl), [&](interpreter & interp) { return interp.run_init(TO_REF(name, decl), TO_REF(name, init_decl)); }); } diff --git a/stage0/src/library/compiler/ir_interpreter.h b/stage0/src/library/compiler/ir_interpreter.h index 00a65a34fe62..c3a0022656df 100644 --- a/stage0/src/library/compiler/ir_interpreter.h +++ b/stage0/src/library/compiler/ir_interpreter.h @@ -5,14 +5,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Sebastian Ullrich */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "runtime/object.h" namespace lean { namespace ir { /** \brief Run `n` using the "boxed" ABI, i.e. with all-owned parameters. */ -object * run_boxed(environment const & env, options const & opts, name const & fn, unsigned n, object **args); -LEAN_EXPORT uint32 run_main(environment const & env, options const & opts, int argv, char * argc[]); +object * run_boxed(elab_environment const & env, options const & opts, name const & fn, unsigned n, object **args); +LEAN_EXPORT uint32 run_main(elab_environment const & env, options const & opts, int argv, char * argc[]); } void initialize_ir_interpreter(); void finalize_ir_interpreter(); diff --git a/stage0/src/library/compiler/lambda_lifting.cpp b/stage0/src/library/compiler/lambda_lifting.cpp index 176db6baae96..e7138092824a 100644 --- a/stage0/src/library/compiler/lambda_lifting.cpp +++ b/stage0/src/library/compiler/lambda_lifting.cpp @@ -30,7 +30,7 @@ bool is_lambda_lifting_name(name fn) { } class lambda_lifting_fn { - environment m_env; + elab_environment m_env; name_generator m_ngen; local_ctx m_lctx; buffer m_new_decls; @@ -39,7 +39,7 @@ class lambda_lifting_fn { typedef std::unordered_set name_set; - environment const & env() { return m_env; } + elab_environment const & env() { return m_env; } name_generator & ngen() { return m_ngen; } @@ -207,10 +207,10 @@ class lambda_lifting_fn { } public: - lambda_lifting_fn(environment const & env): + lambda_lifting_fn(elab_environment const & env): m_env(env) {} - pair operator()(comp_decl const & cdecl) { + pair operator()(comp_decl const & cdecl) { m_base_name = cdecl.fst(); expr r = visit(cdecl.snd(), true); comp_decl new_cdecl(cdecl.fst(), r); @@ -219,11 +219,11 @@ class lambda_lifting_fn { } }; -pair lambda_lifting(environment const & env, comp_decl const & d) { +pair lambda_lifting(elab_environment const & env, comp_decl const & d) { return lambda_lifting_fn(env)(d); } -pair lambda_lifting(environment env, comp_decls const & ds) { +pair lambda_lifting(elab_environment env, comp_decls const & ds) { comp_decls r; for (comp_decl const & d : ds) { comp_decls new_ds; diff --git a/stage0/src/library/compiler/lambda_lifting.h b/stage0/src/library/compiler/lambda_lifting.h index c4f921467c53..ded6caee3772 100644 --- a/stage0/src/library/compiler/lambda_lifting.h +++ b/stage0/src/library/compiler/lambda_lifting.h @@ -5,11 +5,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "library/compiler/util.h" namespace lean { /** \brief Lift lambda expressions in `ds`. The result also contains new auxiliary declarations that have been generated. */ -pair lambda_lifting(environment env, comp_decls const & ds); +pair lambda_lifting(elab_environment env, comp_decls const & ds); /* Return true iff `fn` is the name of an auxiliary function generated by `lambda_lifting`. */ bool is_lambda_lifting_name(name fn); }; diff --git a/stage0/src/library/compiler/lcnf.cpp b/stage0/src/library/compiler/lcnf.cpp index 9bcd88d4cf59..95949e744e88 100644 --- a/stage0/src/library/compiler/lcnf.cpp +++ b/stage0/src/library/compiler/lcnf.cpp @@ -37,6 +37,7 @@ bool is_do_notation_joinpoint(name const & n) { class to_lcnf_fn { typedef rb_expr_map cache; + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; cache m_cache; @@ -44,10 +45,10 @@ class to_lcnf_fn { name m_x; unsigned m_next_idx{1}; public: - to_lcnf_fn(environment const & env, local_ctx const & lctx): - m_st(env), m_lctx(lctx), m_x("_x") {} + to_lcnf_fn(elab_environment const & env, local_ctx const & lctx): + m_env(env), m_st(env), m_lctx(lctx), m_x("_x") {} - environment & env() { return m_st.env(); } + elab_environment & env() { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -637,7 +638,7 @@ class to_lcnf_fn { } }; -expr to_lcnf_core(environment const & env, local_ctx const & lctx, expr const & e) { +expr to_lcnf_core(elab_environment const & env, local_ctx const & lctx, expr const & e) { expr new_e = unfold_macro_defs(env, e); return to_lcnf_fn(env, lctx)(new_e); } diff --git a/stage0/src/library/compiler/lcnf.h b/stage0/src/library/compiler/lcnf.h index 58d3c9c03c60..79fbb4967d70 100644 --- a/stage0/src/library/compiler/lcnf.h +++ b/stage0/src/library/compiler/lcnf.h @@ -5,10 +5,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -expr to_lcnf_core(environment const & env, local_ctx const & lctx, expr const & e); -inline expr to_lcnf(environment const & env, expr const & e) { return to_lcnf_core(env, local_ctx(), e); } +expr to_lcnf_core(elab_environment const & env, local_ctx const & lctx, expr const & e); +inline expr to_lcnf(elab_environment const & env, expr const & e) { return to_lcnf_core(env, local_ctx(), e); } void initialize_lcnf(); void finalize_lcnf(); } diff --git a/stage0/src/library/compiler/ll_infer_type.cpp b/stage0/src/library/compiler/ll_infer_type.cpp index 9167aeb8613b..671c9183c552 100644 --- a/stage0/src/library/compiler/ll_infer_type.cpp +++ b/stage0/src/library/compiler/ll_infer_type.cpp @@ -17,12 +17,13 @@ static expr * g_bot = nullptr; /* Infer type of expressions in ENF or LLNF. */ class ll_infer_type_fn { + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; buffer const * m_new_decl_names{nullptr}; buffer const * m_new_decl_types{nullptr}; - environment const & env() const { return m_st.env(); } + elab_environment const & env() const { return m_env; } name_generator & ngen() { return m_st.ngen(); } bool may_use_bot() const { return m_new_decl_types != nullptr; } @@ -227,14 +228,14 @@ class ll_infer_type_fn { } public: - ll_infer_type_fn(environment const & env, buffer const & ns, buffer const & ts): - m_st(env), m_new_decl_names(&ns), m_new_decl_types(&ts) {} - ll_infer_type_fn(environment const & env, local_ctx const & lctx): - m_st(env), m_lctx(lctx) {} + ll_infer_type_fn(elab_environment const & env, buffer const & ns, buffer const & ts): + m_env(env), m_st(env), m_new_decl_names(&ns), m_new_decl_types(&ts) {} + ll_infer_type_fn(elab_environment const & env, local_ctx const & lctx): + m_env(env), m_st(env), m_lctx(lctx) {} expr operator()(expr const & e) { return infer(e); } }; -void ll_infer_type(environment const & env, comp_decls const & ds, buffer & ts) { +void ll_infer_type(elab_environment const & env, comp_decls const & ds, buffer & ts) { buffer ns; ts.clear(); /* Initialize `ts` */ @@ -272,7 +273,7 @@ void ll_infer_type(environment const & env, comp_decls const & ds, buffer lean_assert(ts.size() == length(ds)); } -expr ll_infer_type(environment const & env, local_ctx const & lctx, expr const & e) { +expr ll_infer_type(elab_environment const & env, local_ctx const & lctx, expr const & e) { return ll_infer_type_fn(env, lctx)(e); } diff --git a/stage0/src/library/compiler/ll_infer_type.h b/stage0/src/library/compiler/ll_infer_type.h index 779f5e958357..35bd795b77fe 100644 --- a/stage0/src/library/compiler/ll_infer_type.h +++ b/stage0/src/library/compiler/ll_infer_type.h @@ -5,11 +5,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -expr ll_infer_type(environment const & env, local_ctx const & lctx, expr const & e); -inline expr ll_infer_type(environment const & env, expr const & e) { return ll_infer_type(env, local_ctx(), e); } -void ll_infer_type(environment const & env, comp_decls const & ds, buffer & ts); +expr ll_infer_type(elab_environment const & env, local_ctx const & lctx, expr const & e); +inline expr ll_infer_type(elab_environment const & env, expr const & e) { return ll_infer_type(env, local_ctx(), e); } +void ll_infer_type(elab_environment const & env, comp_decls const & ds, buffer & ts); void initialize_ll_infer_type(); void finalize_ll_infer_type(); } diff --git a/stage0/src/library/compiler/llnf.cpp b/stage0/src/library/compiler/llnf.cpp index 25a4b4f11b4c..d5af9bbb19f7 100644 --- a/stage0/src/library/compiler/llnf.cpp +++ b/stage0/src/library/compiler/llnf.cpp @@ -247,7 +247,7 @@ cnstr_info::cnstr_info(unsigned cidx, list const & finfo): } } -unsigned get_llnf_arity(environment const & env, name const & n) { +unsigned get_llnf_arity(elab_environment const & env, name const & n) { /* First, try to infer arity from `_cstage2` auxiliary definition. */ name c = mk_cstage2_name(n); optional info = env.find(c); @@ -334,6 +334,7 @@ class to_lambda_pure_fn { typedef name_hash_set name_set; typedef name_hash_map cnstr_info_cache; typedef name_hash_map> enum_cache; + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; buffer m_fvars; @@ -343,7 +344,7 @@ class to_lambda_pure_fn { unsigned m_next_jp_idx{1}; cnstr_info_cache m_cnstr_info_cache; - environment const & env() const { return m_st.env(); } + elab_environment const & env() const { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -803,8 +804,8 @@ class to_lambda_pure_fn { } public: - to_lambda_pure_fn(environment const & env): - m_st(env), m_x("_x"), m_j("j") { + to_lambda_pure_fn(elab_environment const & env): + m_env(env), m_st(env), m_x("_x"), m_j("j") { } expr operator()(expr e) { @@ -815,7 +816,7 @@ class to_lambda_pure_fn { } }; -expr get_constant_ll_type(environment const & env, name const & c) { +expr get_constant_ll_type(elab_environment const & env, name const & c) { if (optional type = get_extern_constant_ll_type(env, c)) { return *type; } else { @@ -823,7 +824,7 @@ expr get_constant_ll_type(environment const & env, name const & c) { } } -environment compile_ir(environment const & env, options const & opts, comp_decls const & ds) { +elab_environment compile_ir(elab_environment const & env, options const & opts, comp_decls const & ds) { buffer new_ds; for (comp_decl const & d : ds) { expr new_v = to_lambda_pure_fn(env)(d.snd()); diff --git a/stage0/src/library/compiler/llnf.h b/stage0/src/library/compiler/llnf.h index 87c61c6ffee5..3b6bebadc1df 100644 --- a/stage0/src/library/compiler/llnf.h +++ b/stage0/src/library/compiler/llnf.h @@ -5,11 +5,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "library/compiler/util.h" namespace lean { -environment compile_ir(environment const & env, options const & opts, comp_decls const & ds); +elab_environment compile_ir(elab_environment const & env, options const & opts, comp_decls const & ds); bool is_llnf_apply(expr const & e); bool is_llnf_closure(expr const & e); @@ -49,8 +49,8 @@ inline bool is_llnf_uset(expr const & e) { unsigned d; return is_llnf_uset(e, d) inline bool is_llnf_box(expr const & e) { unsigned n; return is_llnf_box(e, n); } inline bool is_llnf_unbox(expr const & e) { unsigned n; return is_llnf_unbox(e, n); } -expr get_constant_ll_type(environment const & env, name const & c); -unsigned get_llnf_arity(environment const & env, name const & c); +expr get_constant_ll_type(elab_environment const & env, name const & c); +unsigned get_llnf_arity(elab_environment const & env, name const & c); struct field_info { /* Remark: the position of a scalar value in diff --git a/stage0/src/library/compiler/reduce_arity.cpp b/stage0/src/library/compiler/reduce_arity.cpp index 22e32381867b..cd5135e5d2b7 100644 --- a/stage0/src/library/compiler/reduce_arity.cpp +++ b/stage0/src/library/compiler/reduce_arity.cpp @@ -29,7 +29,7 @@ bool arity_was_reduced(comp_decl const & cdecl) { return is_reduce_arity_aux_fn(n) && n.get_prefix() == cdecl.fst(); } -comp_decls reduce_arity(environment const & env, comp_decl const & cdecl) { +comp_decls reduce_arity(elab_environment const & env, comp_decl const & cdecl) { if (has_export_name(env, cdecl.fst()) || cdecl.fst() == "main") { /* We do not modify the arity of entry points (i.e., functions with attribute [export]) */ return comp_decls(cdecl); @@ -92,7 +92,7 @@ comp_decls reduce_arity(environment const & env, comp_decl const & cdecl) { return comp_decls(red_decl, comp_decls(new_decl)); } -comp_decls reduce_arity(environment const & env, comp_decls const & ds) { +comp_decls reduce_arity(elab_environment const & env, comp_decls const & ds) { comp_decls r; for (comp_decl const & d : ds) { r = append(r, reduce_arity(env, d)); diff --git a/stage0/src/library/compiler/reduce_arity.h b/stage0/src/library/compiler/reduce_arity.h index f16a00bb5852..9c131c75dc2d 100644 --- a/stage0/src/library/compiler/reduce_arity.h +++ b/stage0/src/library/compiler/reduce_arity.h @@ -6,7 +6,7 @@ Author: Leonardo de Moura #pragma once #include "library/compiler/util.h" namespace lean { -comp_decls reduce_arity(environment const & env, comp_decls const & cdecls); +comp_decls reduce_arity(elab_environment const & env, comp_decls const & cdecls); /* Return true if the `cdecl` is of the form `f := fun xs, f._rarg ...`. That is, `f`s arity "was reduced" and an auxiliary declaration `f._rarg` was created to replace it. */ bool arity_was_reduced(comp_decl const & cdecl); diff --git a/stage0/src/library/compiler/simp_app_args.cpp b/stage0/src/library/compiler/simp_app_args.cpp index 84f34ad9c82c..7c18d2d179f8 100644 --- a/stage0/src/library/compiler/simp_app_args.cpp +++ b/stage0/src/library/compiler/simp_app_args.cpp @@ -11,13 +11,14 @@ Author: Leonardo de Moura namespace lean { /* Make sure every argument of applications and projections is a free variable (or neutral element). */ class simp_app_args_fn { + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; buffer m_fvars; name m_x; unsigned m_next_idx{1}; - environment const & env() const { return m_st.env(); } + elab_environment const & env() const { return m_env; } name_generator & ngen() { return m_st.ngen(); } name next_name() { @@ -125,14 +126,14 @@ class simp_app_args_fn { } public: - simp_app_args_fn(environment const & env):m_st(env), m_x("_x") {} + simp_app_args_fn(elab_environment const & env):m_env(env), m_st(env), m_x("_x") {} expr operator()(expr const & e) { return mk_let(0, visit(e)); } }; -expr simp_app_args(environment const & env, expr const & e) { +expr simp_app_args(elab_environment const & env, expr const & e) { return simp_app_args_fn(env)(e); } } diff --git a/stage0/src/library/compiler/simp_app_args.h b/stage0/src/library/compiler/simp_app_args.h index af9bb6e1d742..35dab66e2458 100644 --- a/stage0/src/library/compiler/simp_app_args.h +++ b/stage0/src/library/compiler/simp_app_args.h @@ -5,7 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { -expr simp_app_args(environment const & env, expr const & e); +expr simp_app_args(elab_environment const & env, expr const & e); } diff --git a/stage0/src/library/compiler/specialize.cpp b/stage0/src/library/compiler/specialize.cpp index 6874412c894e..6340fabf7b2d 100644 --- a/stage0/src/library/compiler/specialize.cpp +++ b/stage0/src/library/compiler/specialize.cpp @@ -19,11 +19,11 @@ namespace lean { extern "C" uint8 lean_has_specialize_attribute(object* env, object* n); extern "C" uint8 lean_has_nospecialize_attribute(object* env, object* n); -bool has_specialize_attribute(environment const & env, name const & n) { +bool has_specialize_attribute(elab_environment const & env, name const & n) { return lean_has_specialize_attribute(env.to_obj_arg(), n.to_obj_arg()); } -bool has_nospecialize_attribute(environment const & env, name const & n) { +bool has_nospecialize_attribute(elab_environment const & env, name const & n) { return lean_has_nospecialize_attribute(env.to_obj_arg(), n.to_obj_arg()); } @@ -103,11 +103,11 @@ class spec_info : public object_ref { extern "C" object* lean_add_specialization_info(object* env, object* fn, object* info); extern "C" object* lean_get_specialization_info(object* env, object* fn); -static environment save_specialization_info(environment const & env, name const & fn, spec_info const & si) { - return environment(lean_add_specialization_info(env.to_obj_arg(), fn.to_obj_arg(), si.to_obj_arg())); +static elab_environment save_specialization_info(elab_environment const & env, name const & fn, spec_info const & si) { + return elab_environment(lean_add_specialization_info(env.to_obj_arg(), fn.to_obj_arg(), si.to_obj_arg())); } -static optional get_specialization_info(environment const & env, name const & fn) { +static optional get_specialization_info(elab_environment const & env, name const & fn) { return to_optional(lean_get_specialization_info(env.to_obj_arg(), fn.to_obj_arg())); } @@ -119,7 +119,7 @@ typedef buffer>> spec_info_buffer; Remark: we only create free variables for the header of each declaration. Then, we assume an argument of a recursive call is fixed iff it is a free variable (see `update_spec_info`). */ -static void update_info_buffer(environment const & env, expr e, name_set const & S, spec_info_buffer & info_buffer) { +static void update_info_buffer(elab_environment const & env, expr e, name_set const & S, spec_info_buffer & info_buffer) { while (true) { switch (e.kind()) { case expr_kind::Lambda: @@ -162,7 +162,7 @@ static void update_info_buffer(environment const & env, expr e, name_set const & } } -bool is_class(environment const & env, expr type) { +bool is_class(elab_environment const & env, expr type) { // This is a temporary hack. We do not unfold `type` here, but we should. We will fix it when we reimplement the compiler in Lean. while (is_pi(type)) { type = binding_body(type); @@ -171,7 +171,7 @@ bool is_class(environment const & env, expr type) { return is_constant(type) && is_class(env, const_name(type)); } -environment update_spec_info(environment const & env, comp_decls const & ds) { +elab_environment update_spec_info(elab_environment const & env, comp_decls const & ds) { name_set S; spec_info_buffer d_infos; name_generator ngen; @@ -237,7 +237,7 @@ environment update_spec_info(environment const & env, comp_decls const & ds) { update_info_buffer(env, code, S, d_infos); } /* Update extension */ - environment new_env = env; + elab_environment new_env = env; names mutual_decls = map2(ds, [&](comp_decl const & d) { return d.fst(); }); for (pair> const & info : d_infos) { name const & n = info.first; @@ -255,15 +255,16 @@ environment update_spec_info(environment const & env, comp_decls const & ds) { extern "C" object* lean_cache_specialization(object* env, object* e, object* fn); extern "C" object* lean_get_cached_specialization(object* env, object* e); -static environment cache_specialization(environment const & env, expr const & k, name const & fn) { - return environment(lean_cache_specialization(env.to_obj_arg(), k.to_obj_arg(), fn.to_obj_arg())); +static elab_environment cache_specialization(elab_environment const & env, expr const & k, name const & fn) { + return elab_environment(lean_cache_specialization(env.to_obj_arg(), k.to_obj_arg(), fn.to_obj_arg())); } -static optional get_cached_specialization(environment const & env, expr const & e) { +static optional get_cached_specialization(elab_environment const & env, expr const & e) { return to_optional(lean_get_cached_specialization(env.to_obj_arg(), e.to_obj_arg())); } class specialize_fn { + elab_environment m_env; type_checker::state m_st; csimp_cfg m_cfg; local_ctx m_lctx; @@ -274,7 +275,7 @@ class specialize_fn { unsigned m_next_idx{1}; name_set m_to_respecialize; - environment const & env() { return m_st.env(); } + elab_environment const & env() { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -969,7 +970,8 @@ class specialize_fn { try { expr type = cheap_beta_reduce(type_checker(m_st).infer(code)); declaration aux_ax = mk_axiom(n, names(), type, true /* meta */); - m_st.env() = env().add(aux_ax, false); + m_env = m_env.add(aux_ax, false); + m_st.env() = m_env; } catch (exception &) { /* We may fail to infer the type of code, since it may be recursive This is a workaround. When we re-implement the compiler in Lean, @@ -1149,7 +1151,7 @@ class specialize_fn { for (comp_decl const & new_decl : new_decls) { m_to_respecialize.insert(new_decl.fst()); } - m_st.env() = update_spec_info(env(), new_decls); + m_env = update_spec_info(env(), new_decls); } /* It is only safe to cache when `m_params.size == 0`. See comment above. */ @@ -1162,7 +1164,7 @@ class specialize_fn { i++; } tout() << ">> key: " << trace_pp_expr(key) << "\n";); - m_st.env() = cache_specialization(env(), key, *new_fn_name); + m_env = cache_specialization(env(), key, *new_fn_name); } } expr r = mk_constant(*new_fn_name); @@ -1213,10 +1215,10 @@ class specialize_fn { } public: - specialize_fn(environment const & env, csimp_cfg const & cfg): - m_st(env), m_cfg(cfg), m_at("_at"), m_spec("_spec") {} + specialize_fn(elab_environment const & env, csimp_cfg const & cfg): + m_env(env), m_st(env), m_cfg(cfg), m_at("_at"), m_spec("_spec") {} - pair operator()(comp_decl const & d) { + pair operator()(comp_decl const & d) { m_base_name = d.fst(); lean_trace(name({"compiler", "specialize"}), tout() << "INPUT: " << d.fst() << "\n" << trace_pp_expr(d.snd()) << "\n";); expr new_v = visit(d.snd()); @@ -1225,11 +1227,11 @@ class specialize_fn { } }; -pair specialize_core(environment const & env, comp_decl const & d, csimp_cfg const & cfg) { +pair specialize_core(elab_environment const & env, comp_decl const & d, csimp_cfg const & cfg) { return specialize_fn(env, cfg)(d); } -pair specialize(environment env, comp_decls const & ds, csimp_cfg const & cfg) { +pair specialize(elab_environment env, comp_decls const & ds, csimp_cfg const & cfg) { env = update_spec_info(env, ds); comp_decls r; for (comp_decl const & d : ds) { diff --git a/stage0/src/library/compiler/specialize.h b/stage0/src/library/compiler/specialize.h index a1651e39c967..63061244d1a6 100644 --- a/stage0/src/library/compiler/specialize.h +++ b/stage0/src/library/compiler/specialize.h @@ -5,11 +5,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "library/compiler/util.h" #include "library/compiler/csimp.h" namespace lean { -pair specialize(environment env, comp_decls const & ds, csimp_cfg const & cfg); +pair specialize(elab_environment env, comp_decls const & ds, csimp_cfg const & cfg); void initialize_specialize(); void finalize_specialize(); } diff --git a/stage0/src/library/compiler/struct_cases_on.cpp b/stage0/src/library/compiler/struct_cases_on.cpp index 50d26c874993..9f28c17cc2f0 100644 --- a/stage0/src/library/compiler/struct_cases_on.cpp +++ b/stage0/src/library/compiler/struct_cases_on.cpp @@ -15,6 +15,7 @@ Author: Leonardo de Moura namespace lean { class struct_cases_on_fn { + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; name_set m_scrutinies; /* Set of variables `x` such that there is `casesOn x ...` in the context */ @@ -23,7 +24,7 @@ class struct_cases_on_fn { name m_fld{"_d"}; unsigned m_next_idx{1}; - environment const & env() { return m_st.env(); } + elab_environment const & env() { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -203,8 +204,8 @@ class struct_cases_on_fn { } public: - struct_cases_on_fn(environment const & env): - m_st(env) { + struct_cases_on_fn(elab_environment const & env): + m_env(env), m_st(env) { } expr operator()(expr const & e) { @@ -212,7 +213,7 @@ class struct_cases_on_fn { } }; -expr struct_cases_on(environment const & env, expr const & e) { +expr struct_cases_on(elab_environment const & env, expr const & e) { return struct_cases_on_fn(env)(e); } } diff --git a/stage0/src/library/compiler/struct_cases_on.h b/stage0/src/library/compiler/struct_cases_on.h index 5b534d95e34b..de18defb53ec 100644 --- a/stage0/src/library/compiler/struct_cases_on.h +++ b/stage0/src/library/compiler/struct_cases_on.h @@ -5,7 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { /* Insert `S.casesOn` applications for a structure `S` when @@ -50,5 +50,5 @@ namespace lean { The missing definitional equalities is problematic. For example, the whole algebraic hierarchy in Lean relies on them. */ -expr struct_cases_on(environment const & env, expr const & e); +expr struct_cases_on(elab_environment const & env, expr const & e); } diff --git a/stage0/src/library/compiler/util.cpp b/stage0/src/library/compiler/util.cpp index 96f8ed4c4da7..5d6fa1e37476 100644 --- a/stage0/src/library/compiler/util.cpp +++ b/stage0/src/library/compiler/util.cpp @@ -94,13 +94,13 @@ extern "C" uint8 lean_has_inline_if_reduce_attribute(object* env, object* n); extern "C" uint8 lean_has_macro_inline_attribute(object* env, object* n); extern "C" uint8 lean_has_noinline_attribute(object* env, object* n); -bool has_inline_attribute(environment const & env, name const & n) { return lean_has_inline_attribute(env.to_obj_arg(), n.to_obj_arg()); } -bool has_inline_if_reduce_attribute(environment const & env, name const & n) { return lean_has_inline_if_reduce_attribute(env.to_obj_arg(), n.to_obj_arg()); } -bool has_macro_inline_attribute(environment const & env, name const & n) { return lean_has_macro_inline_attribute(env.to_obj_arg(), n.to_obj_arg()); } -bool has_noinline_attribute(environment const & env, name const & n) { return lean_has_noinline_attribute(env.to_obj_arg(), n.to_obj_arg()); } +bool has_inline_attribute(elab_environment const & env, name const & n) { return lean_has_inline_attribute(env.to_obj_arg(), n.to_obj_arg()); } +bool has_inline_if_reduce_attribute(elab_environment const & env, name const & n) { return lean_has_inline_if_reduce_attribute(env.to_obj_arg(), n.to_obj_arg()); } +bool has_macro_inline_attribute(elab_environment const & env, name const & n) { return lean_has_macro_inline_attribute(env.to_obj_arg(), n.to_obj_arg()); } +bool has_noinline_attribute(elab_environment const & env, name const & n) { return lean_has_noinline_attribute(env.to_obj_arg(), n.to_obj_arg()); } extern "C" uint8 lean_has_never_extract_attribute(object* env, object *n); -bool has_never_extract_attribute(environment const & env, name const & n) { return lean_has_never_extract_attribute(env.to_obj_arg(), n.to_obj_arg()); } +bool has_never_extract_attribute(elab_environment const & env, name const & n) { return lean_has_never_extract_attribute(env.to_obj_arg(), n.to_obj_arg()); } bool is_lcnf_atom(expr const & e) { switch (e.kind()) { @@ -126,8 +126,8 @@ expr elim_trivial_let_decls(expr const & e) { } struct unfold_macro_defs_fn : public replace_visitor { - environment const & m_env; - unfold_macro_defs_fn(environment const & env):m_env(env) {} + elab_environment const & m_env; + unfold_macro_defs_fn(elab_environment const & env):m_env(env) {} bool should_macro_inline(name const & n) { @@ -175,15 +175,15 @@ struct unfold_macro_defs_fn : public replace_visitor { } }; -expr unfold_macro_defs(environment const & env, expr const & e) { +expr unfold_macro_defs(elab_environment const & env, expr const & e) { return unfold_macro_defs_fn(env)(e); } -bool is_cases_on_recursor(environment const & env, name const & n) { +bool is_cases_on_recursor(elab_environment const & env, name const & n) { return ::lean::is_aux_recursor(env, n) && n.get_string() == g_cases_on; } -unsigned get_cases_on_arity(environment const & env, name const & c, bool before_erasure) { +unsigned get_cases_on_arity(elab_environment const & env, name const & c, bool before_erasure) { lean_assert(is_cases_on_recursor(env, c)); inductive_val I_val = get_cases_on_inductive_val(env, c); unsigned nminors = I_val.get_ncnstrs(); @@ -196,7 +196,7 @@ unsigned get_cases_on_arity(environment const & env, name const & c, bool before } } -unsigned get_cases_on_major_idx(environment const & env, name const & c, bool before_erasure) { +unsigned get_cases_on_major_idx(elab_environment const & env, name const & c, bool before_erasure) { if (before_erasure) { inductive_val I_val = get_cases_on_inductive_val(env, c); return I_val.get_nparams() + 1 /* motive */ + I_val.get_nindices(); @@ -205,14 +205,14 @@ unsigned get_cases_on_major_idx(environment const & env, name const & c, bool be } } -expr get_cases_on_app_major(environment const & env, expr const & c, bool before_erasure) { +expr get_cases_on_app_major(elab_environment const & env, expr const & c, bool before_erasure) { lean_assert(is_cases_on_app(env, c)); buffer args; expr const & fn = get_app_args(c, args); return args[get_cases_on_major_idx(env, const_name(fn), before_erasure)]; } -pair get_cases_on_minors_range(environment const & env, name const & c, bool before_erasure) { +pair get_cases_on_minors_range(elab_environment const & env, name const & c, bool before_erasure) { inductive_val I_val = get_cases_on_inductive_val(env, c); unsigned nminors = I_val.get_ncnstrs(); if (before_erasure) { @@ -294,7 +294,7 @@ void sort_fvars(local_ctx const & lctx, buffer & fvars) { }); } -unsigned get_lcnf_size(environment const & env, expr e) { +unsigned get_lcnf_size(elab_environment const & env, expr e) { unsigned r = 0; switch (e.kind()) { case expr_kind::BVar: case expr_kind::MVar: @@ -558,16 +558,16 @@ expr mk_runtime_type(type_checker::state & st, local_ctx const & lctx, expr e) { } } -environment register_stage1_decl(environment const & env, name const & n, names const & ls, expr const & t, expr const & v) { +elab_environment register_stage1_decl(elab_environment const & env, name const & n, names const & ls, expr const & t, expr const & v) { declaration aux_decl = mk_definition(mk_cstage1_name(n), ls, t, v, reducibility_hints::mk_opaque(), definition_safety::unsafe); return env.add(aux_decl, false); } -bool is_stage2_decl(environment const & env, name const & n) { +bool is_stage2_decl(elab_environment const & env, name const & n) { return static_cast(env.find(mk_cstage2_name(n))); } -environment register_stage2_decl(environment const & env, name const & n, expr const & t, expr const & v) { +elab_environment register_stage2_decl(elab_environment const & env, name const & n, expr const & t, expr const & v) { declaration aux_decl = mk_definition(mk_cstage2_name(n), names(), t, v, reducibility_hints::mk_opaque(), definition_safety::unsafe); return env.add(aux_decl, false); @@ -608,10 +608,11 @@ we can get an irrelevant `ty`. We disabled this validator since we will delete the code generator written in C++. */ class lcnf_valid_let_decls_fn { + elab_environment m_env; type_checker::state m_st; local_ctx m_lctx; - environment const & env() const { return m_st.env(); } + elab_environment const & env() const { return m_env; } name_generator & ngen() { return m_st.ngen(); } @@ -675,19 +676,19 @@ class lcnf_valid_let_decls_fn { } public: - lcnf_valid_let_decls_fn(environment const & env, local_ctx const & lctx): - m_st(env), m_lctx(lctx) {} + lcnf_valid_let_decls_fn(elab_environment const & env, local_ctx const & lctx): + m_env(env), m_st(env), m_lctx(lctx) {} optional operator()(expr const & e) { return visit(e); } }; -optional lcnf_valid_let_decls(environment const & env, expr const & e) { +optional lcnf_valid_let_decls(elab_environment const & env, expr const & e) { return lcnf_valid_let_decls_fn(env, local_ctx())(e); } -bool lcnf_check_let_decls(environment const & env, comp_decl const & d) { +bool lcnf_check_let_decls(elab_environment const & env, comp_decl const & d) { if (optional v = lcnf_valid_let_decls(env, d.snd())) { tout() << "LCNF violation at " << d.fst() << "\n" << *v << "\n"; return false; @@ -696,7 +697,7 @@ bool lcnf_check_let_decls(environment const & env, comp_decl const & d) { } } -bool lcnf_check_let_decls(environment const & env, comp_decls const & ds) { +bool lcnf_check_let_decls(elab_environment const & env, comp_decls const & ds) { for (comp_decl const & d : ds) { if (!lcnf_check_let_decls(env, d)) return false; @@ -771,12 +772,12 @@ expr lcnf_eta_expand(type_checker::state & st, local_ctx lctx, expr e) { } } -bool is_quot_primitive_app(environment const & env, expr const & e) { +bool is_quot_primitive_app(elab_environment const & env, expr const & e) { expr const & f = get_app_fn(e); return is_constant(f) && is_quot_primitive(env, const_name(f)); } -bool must_be_eta_expanded(environment const & env, expr const & e) { +bool must_be_eta_expanded(elab_environment const & env, expr const & e) { return is_constructor_app(env, e) || is_proj(e) || diff --git a/stage0/src/library/compiler/util.h b/stage0/src/library/compiler/util.h index bb7e45788403..49a3844e38f0 100644 --- a/stage0/src/library/compiler/util.h +++ b/stage0/src/library/compiler/util.h @@ -12,6 +12,7 @@ Author: Leonardo de Moura #include "kernel/type_checker.h" #include "library/constants.h" #include "library/util.h" +#include "library/elab_environment.h" namespace lean { /* Return the `some(n)` if `I` is the name of an inductive datatype that contains only constructors with 0-arguments, @@ -35,19 +36,19 @@ bool is_lcnf_atom(expr const & e); expr elim_trivial_let_decls(expr const & e); -bool has_inline_attribute(environment const & env, name const & n); -bool has_noinline_attribute(environment const & env, name const & n); -bool has_inline_if_reduce_attribute(environment const & env, name const & n); -bool has_never_extract_attribute(environment const & env, name const & n); +bool has_inline_attribute(elab_environment const & env, name const & n); +bool has_noinline_attribute(elab_environment const & env, name const & n); +bool has_inline_if_reduce_attribute(elab_environment const & env, name const & n); +bool has_never_extract_attribute(elab_environment const & env, name const & n); -expr unfold_macro_defs(environment const & env, expr const & e); +expr unfold_macro_defs(elab_environment const & env, expr const & e); /* Return true if the given argument is mdata relevant to the compiler Remark: we currently don't keep any metadata in the compiler. */ inline bool is_lc_mdata(expr const &) { return false; } -bool is_cases_on_recursor(environment const & env, name const & n); +bool is_cases_on_recursor(elab_environment const & env, name const & n); /* We defined the "arity" of a cases_on application as the sum: ``` number of inductive parameters + @@ -57,30 +58,30 @@ bool is_cases_on_recursor(environment const & env, name const & n); number of constructors // cases_on has a minor premise for each constructor ``` \pre is_cases_on_recursor(env, c) */ -unsigned get_cases_on_arity(environment const & env, name const & c, bool before_erasure = true); +unsigned get_cases_on_arity(elab_environment const & env, name const & c, bool before_erasure = true); /* Return the `inductive_val` for the cases_on constant `c`. */ -inline inductive_val get_cases_on_inductive_val(environment const & env, name const & c) { +inline inductive_val get_cases_on_inductive_val(elab_environment const & env, name const & c) { lean_assert(is_cases_on_recursor(env, c)); return env.get(c.get_prefix()).to_inductive_val(); } -inline inductive_val get_cases_on_inductive_val(environment const & env, expr const & c) { +inline inductive_val get_cases_on_inductive_val(elab_environment const & env, expr const & c) { lean_assert(is_constant(c)); return get_cases_on_inductive_val(env, const_name(c)); } -inline bool is_cases_on_app(environment const & env, expr const & e) { +inline bool is_cases_on_app(elab_environment const & env, expr const & e) { expr const & fn = get_app_fn(e); return is_constant(fn) && is_cases_on_recursor(env, const_name(fn)); } /* Return the major premise of a cases_on-application. \pre is_cases_on_app(env, c) */ -expr get_cases_on_app_major(environment const & env, expr const & c, bool before_erasure = true); -unsigned get_cases_on_major_idx(environment const & env, name const & c, bool before_erasure = true); +expr get_cases_on_app_major(elab_environment const & env, expr const & c, bool before_erasure = true); +unsigned get_cases_on_major_idx(elab_environment const & env, name const & c, bool before_erasure = true); /* Return the pair `(b, e)` such that `i in [b, e)` is argument `i` in a `c` cases_on application is a minor premise. \pre is_cases_on_recursor(env, c) */ -pair get_cases_on_minors_range(environment const & env, name const & c, bool before_erasure = true); +pair get_cases_on_minors_range(elab_environment const & env, name const & c, bool before_erasure = true); -inline bool is_quot_primitive(environment const & env, name const & n) { +inline bool is_quot_primitive(elab_environment const & env, name const & n) { optional info = env.find(n); return info && info->is_quot(); } @@ -119,7 +120,7 @@ expr replace_fvar(expr const & e, expr const & fvar, expr const & new_term); void sort_fvars(local_ctx const & lctx, buffer & fvars); /* Return the "code" size for `e` */ -unsigned get_lcnf_size(environment const & env, expr e); +unsigned get_lcnf_size(elab_environment const & env, expr e); // ======================================= // Auxiliary expressions for erasure. @@ -165,9 +166,9 @@ void collect_used(expr const & e, name_hash_set & S); /* Return true iff `e` contains a free variable in `s` */ bool depends_on(expr const & e, name_hash_set const & s); -bool is_stage2_decl(environment const & env, name const & n); -environment register_stage1_decl(environment const & env, name const & n, names const & ls, expr const & t, expr const & v); -environment register_stage2_decl(environment const & env, name const & n, expr const & t, expr const & v); +bool is_stage2_decl(elab_environment const & env, name const & n); +elab_environment register_stage1_decl(elab_environment const & env, name const & n, names const & ls, expr const & t, expr const & v); +elab_environment register_stage2_decl(elab_environment const & env, name const & n, expr const & t, expr const & v); /* Return `some n` iff `e` is of the forms `expr.lit (literal.nat n)` or `uint*.of_nat (expr.lit (literal.nat n))` */ optional get_num_lit_ext(expr const & e); @@ -181,8 +182,8 @@ optional is_fix_core(name const & c); Remark: this function assumes universe levels have already been erased. */ optional mk_enf_fix_core(unsigned n); -bool lcnf_check_let_decls(environment const & env, comp_decl const & d); -bool lcnf_check_let_decls(environment const & env, comp_decls const & ds); +bool lcnf_check_let_decls(elab_environment const & env, comp_decl const & d); +bool lcnf_check_let_decls(elab_environment const & env, comp_decls const & ds); // ======================================= /* Similar to `type_checker::eta_expand`, but preserves LCNF */ @@ -200,11 +201,11 @@ optional is_enum_type(environment const & env, expr const & type); extern "C" uint8 lean_is_matcher(object* env, object* n); -inline bool is_matcher(environment const & env, name const & n) { +inline bool is_matcher(elab_environment const & env, name const & n) { return lean_is_matcher(env.to_obj_arg(), n.to_obj_arg()); } -inline bool is_matcher_app(environment const & env, expr const & e) { +inline bool is_matcher_app(elab_environment const & env, expr const & e) { expr const & f = get_app_fn(e); return is_constant(f) && is_matcher(env, const_name(f)); } @@ -213,7 +214,7 @@ inline bool is_matcher_app(environment const & env, expr const & e) { Return true if the given expression must be in eta-expanded form during compilation. Example: constructors, `casesOn` applications must always be in eta-expanded form. */ -bool must_be_eta_expanded(environment const & env, expr const & e); +bool must_be_eta_expanded(elab_environment const & env, expr const & e); void initialize_compiler_util(); void finalize_compiler_util(); diff --git a/stage0/src/library/constructions/no_confusion.cpp b/stage0/src/library/constructions/no_confusion.cpp index 6710b0b0dd79..8a3470c4d4db 100644 --- a/stage0/src/library/constructions/no_confusion.cpp +++ b/stage0/src/library/constructions/no_confusion.cpp @@ -6,7 +6,7 @@ Author: Leonardo de Moura */ #include "runtime/sstream.h" #include "kernel/kernel_exception.h" -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "kernel/instantiate.h" #include "kernel/abstract.h" #include "kernel/type_checker.h" @@ -22,7 +22,7 @@ static void throw_corrupted(name const & n) { throw exception(sstream() << "error in '" << g_no_confusion << "' generation, '" << n << "' inductive datatype declaration is corrupted"); } -static declaration mk_no_confusion_type(environment const & env, name const & n) { +static declaration mk_no_confusion_type(elab_environment const & env, name const & n) { constant_info ind_info = env.get(n); inductive_val ind_val = ind_info.to_inductive_val(); local_ctx lctx; @@ -122,10 +122,10 @@ static declaration mk_no_confusion_type(environment const & env, name const & n) } extern "C" LEAN_EXPORT object * lean_mk_no_confusion_type(object * env, object * n) { - return catch_kernel_exceptions([&]() { return mk_no_confusion_type(environment(env), name(n, true)); }); + return catch_kernel_exceptions([&]() { return mk_no_confusion_type(elab_environment(env), name(n, true)); }); } -declaration mk_no_confusion(environment const & env, name const & n) { +declaration mk_no_confusion(elab_environment const & env, name const & n) { local_ctx lctx; name_generator ngen = mk_constructions_name_generator(); constant_info ind_info = env.get(n); @@ -229,6 +229,6 @@ declaration mk_no_confusion(environment const & env, name const & n) { } extern "C" LEAN_EXPORT object * lean_mk_no_confusion(object * env, object * n) { - return catch_kernel_exceptions([&]() { return mk_no_confusion(environment(env), name(n, true)); }); + return catch_kernel_exceptions([&]() { return mk_no_confusion(elab_environment(env), name(n, true)); }); } } diff --git a/stage0/src/library/constructions/no_confusion.h b/stage0/src/library/constructions/no_confusion.h index b15317c90a24..951959c7fb02 100644 --- a/stage0/src/library/constructions/no_confusion.h +++ b/stage0/src/library/constructions/no_confusion.h @@ -5,7 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { /** \brief Given an inductive datatype \c n (which is not a proposition) in \c env, @@ -15,7 +15,7 @@ namespace lean { If the environment has an impredicative Prop, it also assumes heq is defined. If the environment does not have an impredicative Prop, then it also assumes lift is defined. */ -declaration mk_no_confusion_type(environment const & env, name const & n); +declaration mk_no_confusion_type(elab_environment const & env, name const & n); /** \brief Given an inductive datatype \c n (which is not a proposition) in \c env, returns the declaration for n.no_confusion. @@ -24,5 +24,5 @@ declaration mk_no_confusion_type(environment const & env, name const & n); If the environment has an impredicative Prop, it also assumes heq is defined. If the environment does not have an impredicative Prop, then it also assumes lift is defined. */ -declaration mk_no_confusion(environment const & env, name const & n); +declaration mk_no_confusion(elab_environment const & env, name const & n); } diff --git a/stage0/src/library/constructions/projection.cpp b/stage0/src/library/constructions/projection.cpp index a234c172fdb3..0e40d4b6ce05 100644 --- a/stage0/src/library/constructions/projection.cpp +++ b/stage0/src/library/constructions/projection.cpp @@ -30,7 +30,7 @@ static bool is_prop(expr type) { return is_sort(type) && is_zero(sort_level(type)); } -environment mk_projections(environment const & env, name const & n, buffer const & proj_names, bool inst_implicit) { +elab_environment mk_projections(elab_environment const & env, name const & n, buffer const & proj_names, bool inst_implicit) { local_ctx lctx; name_generator ngen = mk_constructions_name_generator(); constant_info ind_info = env.get(n); @@ -82,7 +82,7 @@ environment mk_projections(environment const & env, name const & n, buffer it = instantiate(binding_body(it), local); } unsigned i = 0; - environment new_env = env; + elab_environment new_env = env; for (name const & proj_name : proj_names) { if (!is_pi(cnstr_type)) throw exception(sstream() << "generating projection '" << proj_name << "', '" @@ -126,14 +126,14 @@ environment mk_projections(environment const & env, name const & n, buffer extern "C" LEAN_EXPORT object * lean_mk_projections(object * env, object * struct_name, object * proj_infos, uint8 inst_implicit) { - environment new_env(env); + elab_environment new_env(env); name n(struct_name); list_ref ps(proj_infos); buffer proj_names; for (auto p : ps) { proj_names.push_back(p); } - return catch_kernel_exceptions([&]() { return mk_projections(new_env, n, proj_names, inst_implicit != 0); }); + return catch_kernel_exceptions([&]() { return mk_projections(new_env, n, proj_names, inst_implicit != 0); }); } void initialize_def_projection() { diff --git a/stage0/src/library/elab_environment.cpp b/stage0/src/library/elab_environment.cpp new file mode 100644 index 000000000000..537169f4ec51 --- /dev/null +++ b/stage0/src/library/elab_environment.cpp @@ -0,0 +1,74 @@ +/* +Copyright (c) 2024 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Authors: Leonardo de Moura, Sebastian Ullrich +*/ +#include "runtime/interrupt.h" +#include "kernel/type_checker.h" +#include "kernel/kernel_exception.h" +#include "library/elab_environment.h" +#include "library/compiler/ir_interpreter.h" + +namespace lean { +/* updateBaseAfterKernelAdd (env : Environment) (added : Declaration) (base : Kernel.Environment) : + Environment + + Updates an elab environment with a given kernel environment after the declaration `d` has been + added to it. `d` is used to adjust further elab env data such as registering new namespaces. + + NOTE: Ideally this language switching would not be necessary and we could do all this in Lean + only but the old code generator and `mk_projections` still need a C++ `elab_environment::add`. */ +extern "C" obj_res lean_elab_environment_update_base_after_kernel_add(obj_arg env, obj_arg d, obj_arg kenv); + +elab_environment elab_environment::add(declaration const & d, bool check) const { + environment kenv = to_kernel_env().add(d, check); + return elab_environment(lean_elab_environment_update_base_after_kernel_add(this->to_obj_arg(), d.to_obj_arg(), kenv.to_obj_arg())); +} + +extern "C" LEAN_EXPORT object * lean_elab_add_decl(object * env, size_t max_heartbeat, object * decl, + object * opt_cancel_tk) { + scope_max_heartbeat s(max_heartbeat); + scope_cancel_tk s2(is_scalar(opt_cancel_tk) ? nullptr : cnstr_get(opt_cancel_tk, 0)); + return catch_kernel_exceptions([&]() { + return elab_environment(env).add(declaration(decl, true)); + }); +} + +extern "C" LEAN_EXPORT object * lean_elab_add_decl_without_checking(object * env, object * decl) { + return catch_kernel_exceptions([&]() { + return elab_environment(env).add(declaration(decl, true), false); + }); +} + +extern "C" obj_res lean_elab_environment_to_kernel_env(obj_arg); +environment elab_environment::to_kernel_env() const { + return environment(lean_elab_environment_to_kernel_env(to_obj_arg())); +} + +extern "C" obj_res lean_display_stats(obj_arg env, obj_arg w); +void elab_environment::display_stats() const { + dec_ref(lean_display_stats(to_obj_arg(), io_mk_world())); +} + +extern "C" LEAN_EXPORT lean_object * lean_kernel_is_def_eq(lean_object * obj_env, lean_object * lctx, lean_object * a, lean_object * b) { + elab_environment env(obj_env); + return catch_kernel_exceptions([&]() { + return lean_box(type_checker(env.to_kernel_env(), local_ctx(lctx)).is_def_eq(expr(a), expr(b))); + }); +} + +extern "C" LEAN_EXPORT lean_object * lean_kernel_whnf(lean_object * obj_env, lean_object * lctx, lean_object * a) { + elab_environment env(obj_env); + return catch_kernel_exceptions([&]() { + return type_checker(env.to_kernel_env(), local_ctx(lctx)).whnf(expr(a)).steal(); + }); +} + +extern "C" LEAN_EXPORT lean_object * lean_kernel_check(lean_object * obj_env, lean_object * lctx, lean_object * a) { + elab_environment env(obj_env); + return catch_kernel_exceptions([&]() { + return type_checker(env.to_kernel_env(), local_ctx(lctx)).check(expr(a)).steal(); + }); +} +} diff --git a/stage0/src/library/elab_environment.h b/stage0/src/library/elab_environment.h new file mode 100644 index 000000000000..db21ccf13b49 --- /dev/null +++ b/stage0/src/library/elab_environment.h @@ -0,0 +1,43 @@ +/* +Copyright (c) 2024 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Authors: Leonardo de Moura, Sebastian Ullrich +*/ +#pragma once +#include "kernel/environment.h" + +namespace lean { +/* Wrapper for `Lean.Environment` */ +class LEAN_EXPORT elab_environment : public object_ref { +public: + elab_environment(elab_environment const & other):object_ref(other) {} + elab_environment(elab_environment && other):object_ref(other) {} + explicit elab_environment(b_obj_arg o, bool b):object_ref(o, b) {} + explicit elab_environment(obj_arg o):object_ref(o) {} + ~elab_environment() {} + + elab_environment & operator=(elab_environment const & other) { object_ref::operator=(other); return *this; } + elab_environment & operator=(elab_environment && other) { object_ref::operator=(other); return *this; } + + /** \brief Return information for the constant with name \c n (if it is defined in this environment). */ + optional find(name const & n) const { return to_kernel_env().find(n); }; + + /** \brief Return information for the constant with name \c n. Throws and exception if constant declaration does not exist in this environment. */ + constant_info get(name const & n) const { return to_kernel_env().get(n); }; + + /** \brief Extends the current environment with the given declaration */ + elab_environment add(declaration const & d, bool check = true) const; + + /** \brief Pointer equality */ + friend bool is_eqp(elab_environment const & e1, elab_environment const & e2) { + return e1.raw() == e2.raw(); + } + + void display_stats() const; + + environment to_kernel_env() const; + + operator environment() const { return to_kernel_env(); } +}; +} diff --git a/stage0/src/library/module.cpp b/stage0/src/library/module.cpp index dad1df464906..91d65d22db0c 100644 --- a/stage0/src/library/module.cpp +++ b/stage0/src/library/module.cpp @@ -245,7 +245,7 @@ extern "C" LEAN_EXPORT object * lean_read_module_data(object * fname, object *) def writeModule (env : Environment) (fname : String) : IO Unit := */ extern "C" object * lean_write_module(object * env, object * fname, object *); -void write_module(environment const & env, std::string const & olean_fn) { +void write_module(elab_environment const & env, std::string const & olean_fn) { consume_io_result(lean_write_module(env.to_obj_arg(), mk_string(olean_fn), io_mk_world())); } } diff --git a/stage0/src/library/module.h b/stage0/src/library/module.h index ae6e508c0d82..cbf161834725 100644 --- a/stage0/src/library/module.h +++ b/stage0/src/library/module.h @@ -10,9 +10,9 @@ Authors: Leonardo de Moura, Gabriel Ebner, Sebastian Ullrich #include #include #include "runtime/optional.h" -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { /** \brief Store module using \c env. */ -LEAN_EXPORT void write_module(environment const & env, std::string const & olean_fn); +LEAN_EXPORT void write_module(elab_environment const & env, std::string const & olean_fn); } diff --git a/stage0/src/library/projection.cpp b/stage0/src/library/projection.cpp index 62a2319022dd..909b385ad0c2 100644 --- a/stage0/src/library/projection.cpp +++ b/stage0/src/library/projection.cpp @@ -25,11 +25,11 @@ bool projection_info::is_inst_implicit() const { return lean_projection_info_fro extern "C" object* lean_add_projection_info(object* env, object* p, object* ctor, object* nparams, object* i, uint8 fromClass); extern "C" object* lean_get_projection_info(object* env, object* p); -environment save_projection_info(environment const & env, name const & p, name const & mk, unsigned nparams, unsigned i, bool inst_implicit) { - return environment(lean_add_projection_info(env.to_obj_arg(), p.to_obj_arg(), mk.to_obj_arg(), mk_nat_obj(nparams), mk_nat_obj(i), inst_implicit)); +elab_environment save_projection_info(elab_environment const & env, name const & p, name const & mk, unsigned nparams, unsigned i, bool inst_implicit) { + return elab_environment(lean_add_projection_info(env.to_obj_arg(), p.to_obj_arg(), mk.to_obj_arg(), mk_nat_obj(nparams), mk_nat_obj(i), inst_implicit)); } -optional get_projection_info(environment const & env, name const & p) { +optional get_projection_info(elab_environment const & env, name const & p) { return to_optional(lean_get_projection_info(env.to_obj_arg(), p.to_obj_arg())); } } diff --git a/stage0/src/library/projection.h b/stage0/src/library/projection.h index 3ade6827e63c..89e88d500447 100644 --- a/stage0/src/library/projection.h +++ b/stage0/src/library/projection.h @@ -5,7 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #pragma once -#include "kernel/environment.h" +#include "library/elab_environment.h" namespace lean { /** \brief Auxiliary information attached to projections. This information @@ -37,20 +37,20 @@ class projection_info : public object_ref { bool is_inst_implicit() const; }; -/** \brief Mark \c p as a projection in the given environment and store that +/** \brief Mark \c p as a projection in the given elab_environment and store that \c mk is the constructor associated with it, \c nparams is the number of parameters, and \c i says that \c p is the i-th projection. */ -environment save_projection_info(environment const & env, name const & p, name const & mk, unsigned nparams, unsigned i, +elab_environment save_projection_info(elab_environment const & env, name const & p, name const & mk, unsigned nparams, unsigned i, bool inst_implicit); -/** \brief If \c p is a projection in the given environment, then return the information +/** \brief If \c p is a projection in the given elab_environment, then return the information associated with it (constructor, number of parameters, and index). If \c p is not a projection, then return nullptr. */ -optional get_projection_info(environment const & env, name const & p); +optional get_projection_info(elab_environment const & env, name const & p); -inline bool is_projection(environment const & env, name const & n) { +inline bool is_projection(elab_environment const & env, name const & n) { return static_cast(get_projection_info(env, n)); } } diff --git a/stage0/src/library/reducible.cpp b/stage0/src/library/reducible.cpp index 16a0aad85f1a..a2f009f6406c 100644 --- a/stage0/src/library/reducible.cpp +++ b/stage0/src/library/reducible.cpp @@ -5,20 +5,20 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Leonardo de Moura */ #include -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "library/reducible.h" namespace lean { extern "C" uint8 lean_get_reducibility_status(object * env, object * n); extern "C" object * lean_set_reducibility_status(object * env, object * n, uint8 s); -environment set_reducible(environment const & env, name const & n, reducible_status s, bool persistent) { +elab_environment set_reducible(elab_environment const & env, name const & n, reducible_status s, bool persistent) { if (!persistent) throw exception("reducibility attributes must be persistent for now, we will relax this restriction in a near future"); - return environment(lean_set_reducibility_status(env.to_obj_arg(), n.to_obj_arg(), static_cast(s))); + return elab_environment(lean_set_reducibility_status(env.to_obj_arg(), n.to_obj_arg(), static_cast(s))); } -reducible_status get_reducible_status(environment const & env, name const & n) { +reducible_status get_reducible_status(elab_environment const & env, name const & n) { return static_cast(lean_get_reducibility_status(env.to_obj_arg(), n.to_obj_arg())); } } diff --git a/stage0/src/library/reducible.h b/stage0/src/library/reducible.h index c58e210426ce..f1fa9e44d2a7 100644 --- a/stage0/src/library/reducible.h +++ b/stage0/src/library/reducible.h @@ -7,6 +7,7 @@ Author: Leonardo de Moura #pragma once #include #include "library/util.h" +#include "library/elab_environment.h" namespace lean { enum class reducible_status { Reducible, Semireducible, Irreducible }; @@ -20,10 +21,10 @@ enum class reducible_status { Reducible, Semireducible, Irreducible }; "Reducible" definitions can be freely unfolded by automation (i.e., elaborator, simplifier, etc). We should view it as a hint to automation. */ -environment set_reducible(environment const & env, name const & n, reducible_status s, bool persistent); +elab_environment set_reducible(elab_environment const & env, name const & n, reducible_status s, bool persistent); -reducible_status get_reducible_status(environment const & env, name const & n); +reducible_status get_reducible_status(elab_environment const & env, name const & n); -inline bool is_reducible(environment const & env, name const & n) { return get_reducible_status(env, n) == reducible_status::Reducible; } -inline bool is_semireducible(environment const & env, name const & n) { return get_reducible_status(env, n) == reducible_status::Semireducible; } +inline bool is_reducible(elab_environment const & env, name const & n) { return get_reducible_status(env, n) == reducible_status::Reducible; } +inline bool is_semireducible(elab_environment const & env, name const & n) { return get_reducible_status(env, n) == reducible_status::Semireducible; } } diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index a3d118fd9f11..b647d85f376c 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -9,7 +9,7 @@ options get_default_options() { opts = opts.update({"debug", "proofAsSorry"}, false); // switch to `true` for ABI-breaking changes affecting meta code; // see also next option! - opts = opts.update({"interpreter", "prefer_native"}, true); + opts = opts.update({"interpreter", "prefer_native"}, false); // switch to `false` when enabling `prefer_native` should also affect use // of built-in parsers in quotations; this is usually the case, but setting // both to `true` may be necessary for handling non-builtin parsers with diff --git a/stage0/src/util/shell.cpp b/stage0/src/util/shell.cpp index a39ca0b6586c..fbc432623ad4 100644 --- a/stage0/src/util/shell.cpp +++ b/stage0/src/util/shell.cpp @@ -28,7 +28,7 @@ Author: Leonardo de Moura #include "util/io.h" #include "util/options.h" #include "util/option_declarations.h" -#include "kernel/environment.h" +#include "library/elab_environment.h" #include "kernel/kernel_exception.h" #include "kernel/trace.h" #include "library/formatter.h" @@ -335,7 +335,7 @@ extern "C" object * lean_run_frontend( object * error_kinds, object * w ); -pair_ref run_new_frontend( +pair_ref run_new_frontend( std::string const & input, options const & opts, std::string const & file_name, name const & main_module_name, @@ -348,7 +348,7 @@ pair_ref run_new_frontend( if (ilean_file_name) { oilean_file_name = mk_option_some(mk_string(*ilean_file_name)); } - return get_io_result>(lean_run_frontend( + return get_io_result>(lean_run_frontend( mk_string(input), opts.to_obj_arg(), mk_string(file_name), @@ -406,7 +406,7 @@ void print_imports_json(array_ref const & fnames) { } extern "C" object* lean_environment_free_regions(object * env, object * w); -void environment_free_regions(environment && env) { +void environment_free_regions(elab_environment && env) { consume_io_result(lean_environment_free_regions(env.steal(), io_mk_world())); } } @@ -643,7 +643,6 @@ extern "C" LEAN_EXPORT int lean_main(int argc, char ** argv) { report_profiling_time("initialization", init_time); } - environment env(trust_lvl); scoped_task_manager scope_task_man(num_threads); optional main_module_name; @@ -715,8 +714,8 @@ extern "C" LEAN_EXPORT int lean_main(int argc, char ** argv) { if (!main_module_name) main_module_name = name("_stdin"); - pair_ref r = run_new_frontend(contents, opts, mod_fn, *main_module_name, trust_lvl, ilean_fn, json_output, error_kinds); - env = r.fst(); + pair_ref r = run_new_frontend(contents, opts, mod_fn, *main_module_name, trust_lvl, ilean_fn, json_output, error_kinds); + elab_environment env = r.fst(); bool ok = unbox(r.snd().raw()); if (stats) { diff --git a/stage0/stdlib/Init/Data/List/Lemmas.c b/stage0/stdlib/Init/Data/List/Lemmas.c index 82a167b40dcc..65fc6dc69c83 100644 --- a/stage0/stdlib/Init/Data/List/Lemmas.c +++ b/stage0/stdlib/Init/Data/List/Lemmas.c @@ -192,39 +192,60 @@ x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_isEqv return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -if (lean_obj_tag(x_2) == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_5; +lean_object* x_6; +lean_dec(x_5); lean_dec(x_4); -x_5 = lean_apply_1(x_3, x_1); -return x_5; +x_6 = lean_apply_1(x_3, lean_box(0)); +return x_6; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_object* x_7; lean_dec(x_3); -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_1, 1); lean_inc(x_7); -lean_dec(x_2); -x_8 = lean_apply_3(x_4, x_1, x_6, x_7); -return x_8; +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_5); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_apply_2(x_4, x_8, lean_box(0)); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_4); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_7, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_7, 1); +lean_inc(x_12); +lean_dec(x_7); +x_13 = lean_apply_4(x_5, x_10, x_11, x_12, lean_box(0)); +return x_13; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter___rarg), 4, 0); -return x_4; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter___rarg), 5, 0); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -246,24 +267,24 @@ return x_6; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg(x_1, x_2, x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -274,221 +295,178 @@ return x_2; } else { -lean_object* x_4; lean_object* x_5; +lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); lean_dec(x_1); -x_5 = lean_apply_1(x_3, x_4); -return x_5; +x_6 = lean_apply_2(x_3, x_4, x_5); +return x_6; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_List_Lemmas_0__List_getLast_x3f_match__1_splitter___rarg(x_1, x_2, x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_dec(x_2); -lean_inc(x_3); -return x_3; +lean_object* x_5; +lean_dec(x_4); +x_5 = lean_apply_1(x_3, x_2); +return x_5; } else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); +lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); lean_dec(x_1); -x_5 = lean_apply_1(x_2, x_4); -return x_5; +x_8 = lean_apply_3(x_4, x_6, x_7, x_2); +return x_8; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter___rarg), 4, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_List_foldlRecOn___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = lean_apply_4(x_1, x_2, x_3, x_4, lean_box(0)); -return x_6; -} -} -LEAN_EXPORT lean_object* l_List_foldlRecOn___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); -return x_4; +lean_inc(x_2); +return x_2; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); lean_dec(x_1); -lean_inc(x_2); -lean_inc(x_6); -lean_inc(x_3); -x_8 = lean_apply_2(x_2, x_3, x_6); -lean_inc(x_5); -x_9 = lean_apply_4(x_5, x_3, x_4, x_6, lean_box(0)); -x_10 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg___lambda__1), 5, 1); -lean_closure_set(x_10, 0, x_5); -x_1 = x_7; -x_3 = x_8; -x_4 = x_9; -x_5 = x_10; -goto _start; +x_5 = lean_apply_1(x_3, x_4); +return x_5; } } } -LEAN_EXPORT lean_object* l_List_foldlRecOn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg), 5, 0); +x_4 = l___private_Init_Data_List_Lemmas_0__List_filterMap_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_List_foldrRecOn___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_dec(x_5); lean_dec(x_2); -lean_inc(x_4); -return x_4; +lean_inc(x_3); +return x_3; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); lean_dec(x_1); -lean_inc(x_7); -lean_inc(x_2); -x_8 = l_List_foldr___rarg(x_2, x_3, x_7); -lean_inc(x_5); -x_9 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg___lambda__1), 5, 1); -lean_closure_set(x_9, 0, x_5); -x_10 = l_List_foldrRecOn___rarg(x_7, x_2, x_3, x_4, x_9); -x_11 = lean_apply_4(x_5, x_8, x_10, x_6, lean_box(0)); -return x_11; +x_5 = lean_apply_1(x_2, x_4); +return x_5; } } } -LEAN_EXPORT lean_object* l_List_foldrRecOn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_List_foldrRecOn___rarg___boxed), 5, 0); -return x_4; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg___boxed), 3, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_List_foldrRecOn___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; -x_6 = l_List_foldrRecOn___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); +lean_object* x_4; +x_4 = l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg(x_1, x_2, x_3); lean_dec(x_3); -return x_6; +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_6; -lean_dec(x_5); -lean_dec(x_4); -x_6 = lean_apply_1(x_3, lean_box(0)); -return x_6; -} -else -{ -lean_object* x_7; -lean_dec(x_3); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; lean_object* x_9; -lean_dec(x_5); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_dec(x_1); -x_9 = lean_apply_2(x_4, x_8, lean_box(0)); -return x_9; +lean_dec(x_2); +lean_inc(x_3); +return x_3; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_4); -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); lean_dec(x_1); -x_11 = lean_ctor_get(x_7, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_dec(x_7); -x_13 = lean_apply_4(x_5, x_10, x_11, x_12, lean_box(0)); -return x_13; -} +x_5 = lean_apply_1(x_2, x_4); +return x_5; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter___rarg), 5, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -499,67 +477,65 @@ return x_2; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; +lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); lean_dec(x_1); -x_6 = lean_apply_2(x_3, x_4, x_5); -return x_6; +x_5 = lean_apply_1(x_3, x_4); +return x_5; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Init_Data_List_Lemmas_0__List_getLast_x21_match__1_splitter___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg(x_1, x_2, x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -if (lean_obj_tag(x_1) == 0) +if (lean_obj_tag(x_2) == 0) { lean_object* x_5; lean_dec(x_4); -x_5 = lean_apply_1(x_3, x_2); +x_5 = lean_apply_1(x_3, x_1); return x_5; } else { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_dec(x_3); -x_6 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_2, 0); lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_2, 1); lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_apply_3(x_4, x_6, x_7, x_2); +lean_dec(x_2); +x_8 = lean_apply_3(x_4, x_1, x_6, x_7); return x_8; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_getLastD_match__1_splitter___rarg), 4, 0); -return x_3; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter___rarg), 4, 0); +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -579,95 +555,119 @@ return x_5; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg___boxed), 3, 0); +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg___boxed), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Init_Data_List_Lemmas_0__Option_isSome_match__1_splitter___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_List_Lemmas_0__List_foldl__filterMap_match__1_splitter___rarg(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldlRecOn___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = lean_apply_4(x_1, x_2, x_3, x_4, lean_box(0)); +return x_6; +} +} +LEAN_EXPORT lean_object* l_List_foldlRecOn___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { if (lean_obj_tag(x_1) == 0) { +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); -lean_inc(x_3); -return x_3; +return x_4; } else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); lean_dec(x_1); -x_5 = lean_apply_1(x_2, x_4); -return x_5; -} -} +lean_inc(x_2); +lean_inc(x_6); +lean_inc(x_3); +x_8 = lean_apply_2(x_2, x_3, x_6); +lean_inc(x_5); +x_9 = lean_apply_4(x_5, x_3, x_4, x_6, lean_box(0)); +x_10 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg___lambda__1), 5, 1); +lean_closure_set(x_10, 0, x_5); +x_1 = x_7; +x_3 = x_8; +x_4 = x_9; +x_5 = x_10; +goto _start; } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg___boxed), 3, 0); -return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldlRecOn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Init_Data_List_Lemmas_0__List_findSome_x3f_match__1_splitter___rarg(x_1, x_2, x_3); -lean_dec(x_3); +x_4 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg), 5, 0); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldrRecOn___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_dec(x_3); -lean_inc(x_2); -return x_2; +lean_dec(x_5); +lean_dec(x_2); +lean_inc(x_4); +return x_4; } else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); lean_dec(x_1); -x_5 = lean_apply_1(x_3, x_4); -return x_5; +lean_inc(x_7); +lean_inc(x_2); +x_8 = l_List_foldr___rarg(x_2, x_3, x_7); +lean_inc(x_5); +x_9 = lean_alloc_closure((void*)(l_List_foldlRecOn___rarg___lambda__1), 5, 1); +lean_closure_set(x_9, 0, x_5); +x_10 = l_List_foldrRecOn___rarg(x_7, x_2, x_3, x_4, x_9); +x_11 = lean_apply_4(x_5, x_8, x_10, x_6, lean_box(0)); +return x_11; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_foldrRecOn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg___boxed), 3, 0); -return x_3; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_List_foldrRecOn___rarg___boxed), 5, 0); +return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldrRecOn___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_4; -x_4 = l___private_Init_Data_List_Lemmas_0__List_filterMap__replicate_match__1_splitter___rarg(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; +lean_object* x_6; +x_6 = l_List_foldrRecOn___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_6; } } LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_partition_loop_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { diff --git a/stage0/stdlib/Init/Data/Vector/Basic.c b/stage0/stdlib/Init/Data/Vector/Basic.c index 77c6325f56a5..d5bf694601f8 100644 --- a/stage0/stdlib/Init/Data/Vector/Basic.c +++ b/stage0/stdlib/Init/Data/Vector/Basic.c @@ -22,7 +22,6 @@ static lean_object* l_Vector_lex___rarg___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_uget___rarg(lean_object*, size_t, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__2; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9; LEAN_EXPORT lean_object* l_Vector_drop___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__2; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__28; @@ -32,20 +31,20 @@ LEAN_EXPORT uint8_t l_Vector_instBEq___rarg(lean_object*, lean_object*, lean_obj static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__36; uint8_t l_Array_isEqvAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__20; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__19; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__31; LEAN_EXPORT lean_object* l_Vector_map___rarg(lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__9; LEAN_EXPORT lean_object* l_Vector_isEqv(lean_object*); LEAN_EXPORT lean_object* l_Vector_back_x21___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_mkVector___rarg(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8; LEAN_EXPORT lean_object* l_Vector_back_x3f___rarg___boxed(lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__17; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__22; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14; LEAN_EXPORT lean_object* l_Vector_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_any___spec__1(lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__30; @@ -56,17 +55,18 @@ static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__V LEAN_EXPORT lean_object* l_Vector_eraseIdx_x21___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1(lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__37; +LEAN_EXPORT lean_object* l_Vector_foldl(lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__10; LEAN_EXPORT lean_object* l_Vector_set_x21___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24; LEAN_EXPORT lean_object* l_Array_toVector___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_lex___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Vector_map___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Vector_isEqv___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_pop___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_eraseIdx___boxed(lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__14; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18; LEAN_EXPORT lean_object* l_instReprVector___rarg(lean_object*); lean_object* lean_array_fswap(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instReprVector___boxed(lean_object*, lean_object*); @@ -79,6 +79,7 @@ LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Data_Vector_Basic_0__r LEAN_EXPORT lean_object* l_Vector_back___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_back_x21(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_toVector(lean_object*); LEAN_EXPORT lean_object* l_Vector_get(lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -86,17 +87,22 @@ static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__25; LEAN_EXPORT lean_object* l_Vector_swap___boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21; LEAN_EXPORT lean_object* l_Vector_swap___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__19; static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__8; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34; LEAN_EXPORT lean_object* l_Vector_reverse___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1599_; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__13; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19; +LEAN_EXPORT lean_object* l_Vector_foldrM___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_swap(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__7; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldr___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__14; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_set_x21___rarg(lean_object*, lean_object*, lean_object*); @@ -105,28 +111,32 @@ LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Vector_instBEq___spec__1___rarg(lean_o LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_any___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__4; lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__25; LEAN_EXPORT uint8_t l_Vector_contains___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instGetElemNatLt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_all(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31; LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22; LEAN_EXPORT lean_object* l_Vector_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1443_; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__40; LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instLE___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldrM(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__16; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; LEAN_EXPORT lean_object* l_Vector_eraseIdx(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_uget___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20; LEAN_EXPORT lean_object* l_Vector_singleton___rarg(lean_object*); LEAN_EXPORT lean_object* l_Vector_contains(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27; static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__13; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3; LEAN_EXPORT lean_object* l_Vector_all___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Vector_instBEq___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Vector_instLT___boxed(lean_object*, lean_object*, lean_object*); @@ -134,11 +144,12 @@ LEAN_EXPORT lean_object* l_Vector_instGetElemNatLt___boxed(lean_object*, lean_ob static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__14; LEAN_EXPORT lean_object* l_Vector_any___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_indexOf_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1929_; LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_anyM___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5; LEAN_EXPORT lean_object* l_Vector_extract(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_tail___rarg(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4; LEAN_EXPORT lean_object* l_Vector_setIfInBounds(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__34; LEAN_EXPORT lean_object* l_Vector_get___rarg(lean_object*, lean_object*); @@ -148,12 +159,12 @@ static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed_ LEAN_EXPORT lean_object* l_Vector_cast___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_all___spec__1(lean_object*); LEAN_EXPORT uint8_t l_Vector_isPrefixOf___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_mkEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Vector_elimAsArray___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instHAppendHAddNat___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_head___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31; lean_object* l_Array_range___lambda__1___boxed(lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -164,36 +175,37 @@ LEAN_EXPORT lean_object* l_Vector_swapIfInBounds___rarg___boxed(lean_object*, le static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__8; static lean_object* l_Vector_swapAt_x21___rarg___closed__3; LEAN_EXPORT lean_object* l_instReprVector(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldr(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_back(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__8; lean_object* l_Array_reverse___rarg(lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15; LEAN_EXPORT lean_object* l_Vector_getD___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__7; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9; LEAN_EXPORT lean_object* l_Vector_pop(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_uget___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29; LEAN_EXPORT lean_object* l_Vector_extract___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13; LEAN_EXPORT lean_object* l_Vector_allM(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldl___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20; LEAN_EXPORT lean_object* l_Vector_cast___rarg(lean_object*); lean_object* l_Array_zipWithAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldr___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_get___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_back_x3f___boxed(lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT uint8_t l_Vector_all___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_isPrefixOf___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Array_isPrefixOf___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__16; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Vector_map___spec__1___rarg(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Vector_contains___boxed(lean_object*, lean_object*); lean_object* l_Array_ofFn___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_all___boxed(lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__4; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23; LEAN_EXPORT lean_object* l_Vector_cast___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Vector_extract___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__15; @@ -206,10 +218,14 @@ static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed_ LEAN_EXPORT lean_object* l_Vector_head___rarg___boxed(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__9; LEAN_EXPORT lean_object* l_Vector_isPrefixOf___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swap___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__10; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Array_back_x3f___rarg(lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30; LEAN_EXPORT lean_object* l_panic___at_Vector_swapAt_x21___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_array_pop(lean_object*); LEAN_EXPORT lean_object* l_Vector_reverse(lean_object*, lean_object*); @@ -217,17 +233,16 @@ static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed_ static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__27; LEAN_EXPORT lean_object* l_Vector_zipWith___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__9; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17; static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__16; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_map(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_take___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_getD(lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__8; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24; LEAN_EXPORT lean_object* l_Vector_anyM___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_set_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_setIfInBounds___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -235,7 +250,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_all___spec__1___rarg static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__18; static lean_object* l_Vector_lex___rarg___closed__1; LEAN_EXPORT lean_object* l_Vector_back_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1517_; LEAN_EXPORT lean_object* l_Vector_append(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Vector_swapAt_x21___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_lex(lean_object*); @@ -243,6 +257,7 @@ LEAN_EXPORT lean_object* l_Vector_swapAt_x21(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__18; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__20; lean_object* l_Array_anyMUnsafe_any___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Vector_foldrM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__19; LEAN_EXPORT uint8_t l_Vector_lex___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Vector_ofFn___boxed(lean_object*, lean_object*, lean_object*); @@ -253,10 +268,14 @@ static lean_object* l_Vector_eraseIdx_x21___rarg___closed__4; LEAN_EXPORT lean_object* l_Vector_elimAsList___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_get___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldl___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldr___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_swapAt_x21___rarg___closed__1; LEAN_EXPORT lean_object* l_Vector_drop___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352_; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__13; LEAN_EXPORT lean_object* l_Vector_elimAsArray___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_lex___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -265,25 +284,26 @@ lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_eraseIdx_x21(lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_lex___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12; static lean_object* l_Vector_swapAt_x21___rarg___closed__4; LEAN_EXPORT lean_object* l_Vector_back_x3f___rarg(lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__26; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__5; LEAN_EXPORT lean_object* l_Vector_swapAt_x21___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_eraseIdx_x21___rarg___closed__2; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30; LEAN_EXPORT lean_object* l_Vector_set___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__17; lean_object* l_Array_indexOfAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16; +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1591_; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__7; LEAN_EXPORT lean_object* l_Vector_back_x21___rarg___boxed(lean_object*, lean_object*); lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_take___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_any(lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__3; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16; LEAN_EXPORT lean_object* l_Vector_indexOf_x3f___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldlM___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_head___rarg(lean_object*); LEAN_EXPORT lean_object* l_Vector_zipWith___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapAt___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -296,24 +316,25 @@ LEAN_EXPORT lean_object* l_Vector_uget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapAt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapAt_x21___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__6; -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1435_; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Vector_eraseIdx_x21___spec__1___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Vector_instLE(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__decEqVector____x40_Init_Data_Vector_Basic___hyg_98____rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__3; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8; lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg___lambda__1(lean_object*, uint8_t); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__39; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Vector_allM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Vector_extract___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldl___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_lex___rarg___lambda__1___boxed(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__2; LEAN_EXPORT lean_object* l_Vector_getD___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13; LEAN_EXPORT lean_object* l_Vector_instGetElemNatLt___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11; LEAN_EXPORT lean_object* l_Vector_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_push___boxed(lean_object*, lean_object*); @@ -323,26 +344,28 @@ static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed_ static lean_object* l_Vector_eraseIdx_x21___rarg___closed__3; static lean_object* l_Vector_swapAt_x21___rarg___closed__2; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__31; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Vector_instBEq___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_any___rarg___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__22; +LEAN_EXPORT lean_object* l_Vector_foldlM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__29; LEAN_EXPORT lean_object* l_panic___at_Vector_swapAt_x21___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____spec__1(lean_object*); lean_object* lean_string_length(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__21; LEAN_EXPORT lean_object* l_Vector_allM___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196_; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7; uint8_t l_Array_contains___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__23; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Vector_instHAppendHAddNat___closed__1; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__33; +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1673_; LEAN_EXPORT lean_object* l_Vector_isPrefixOf(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapIfInBounds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Vector_foldlM(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__10; static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__6; static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__23; @@ -351,10 +374,10 @@ static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__V LEAN_EXPORT lean_object* l_Array_toVector___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Vector_cast(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__17; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swapAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__21; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3; lean_object* lean_panic_fn(lean_object*, lean_object*); uint8_t l_Array_instDecidableEq___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__4; @@ -365,6 +388,7 @@ lean_object* l_Array_back_x21___rarg(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__38; static lean_object* l_Vector_range___closed__1; LEAN_EXPORT lean_object* l_Vector_append___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__7; LEAN_EXPORT lean_object* l_Vector_tail___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__5; @@ -372,7 +396,7 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Vector_all___spec__1___rarg(lean static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__12; LEAN_EXPORT lean_object* l_Vector_mkEmpty___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Vector_set___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12; +size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__15; LEAN_EXPORT lean_object* l_panic___at_Vector_eraseIdx_x21___spec__1(lean_object*, lean_object*); @@ -385,22 +409,23 @@ LEAN_EXPORT lean_object* l_Vector_head(lean_object*, lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Vector_allM___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__35; LEAN_EXPORT lean_object* l_instDecidableEqVector(lean_object*, lean_object*); +lean_object* l_Array_foldrMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23; LEAN_EXPORT lean_object* l_Vector_back___rarg___boxed(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__27; LEAN_EXPORT lean_object* l_Vector_reverse___rarg(lean_object*); size_t lean_array_size(lean_object*); +lean_object* l_Array_foldlMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_instDecidableEqVector___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__11; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26; LEAN_EXPORT lean_object* l_Vector_instInhabited(lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__4; LEAN_EXPORT lean_object* l_Vector_swapIfInBounds___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__9; LEAN_EXPORT lean_object* l_Vector_swapIfInBounds___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11; -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_set_x21___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_drop(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -410,28 +435,31 @@ LEAN_EXPORT lean_object* l_Vector_zipWith___rarg___boxed(lean_object*, lean_obje static lean_object* l_Vector_eraseIdx_x21___rarg___closed__1; LEAN_EXPORT lean_object* l_Vector_instMembership___boxed(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1773_; LEAN_EXPORT lean_object* l_Vector_indexOf_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_append___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__5; LEAN_EXPORT lean_object* l_Vector_set(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_swap(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_instGetElemNatLt___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Vector_ofFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__1; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__3; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__15; static lean_object* l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__4; +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28; lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32; LEAN_EXPORT lean_object* l_Vector_setIfInBounds___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__18; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_zipWith(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__12; LEAN_EXPORT lean_object* l_Vector_range(lean_object*); +static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17; static lean_object* l_Vector_term_x23v_x5b___x2c_x5d___closed__17; LEAN_EXPORT lean_object* l_Vector_mkEmpty(lean_object*); LEAN_EXPORT lean_object* l_Vector_eraseIdx_x21___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -443,7 +471,6 @@ LEAN_EXPORT lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_1091_; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_Vector_anyM(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2; static lean_object* l___private_Init_Data_Vector_Basic_0__reprVector____x40_Init_Data_Vector_Basic___hyg_32____rarg___closed__13; LEAN_EXPORT lean_object* l_Vector_elimAsList(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Vector_mkVector(lean_object*); @@ -2677,6 +2704,398 @@ lean_dec(x_2); return x_3; } } +LEAN_EXPORT lean_object* l_Vector_foldlM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_apply_2(x_9, lean_box(0), x_3); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = lean_nat_dec_le(x_5, x_5); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_apply_2(x_13, lean_box(0), x_3); +return x_14; +} +else +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = 0; +x_16 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_17 = l_Array_foldlMUnsafe_fold___rarg(x_1, x_2, x_4, x_15, x_16, x_3); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l_Vector_foldlM(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Vector_foldlM___rarg), 4, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Vector_foldlM___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Vector_foldlM(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Vector_foldrM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_apply_2(x_9, lean_box(0), x_3); +return x_10; +} +else +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_12 = 0; +x_13 = l_Array_foldrMUnsafe_fold___rarg(x_1, x_2, x_4, x_11, x_12, x_3); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Vector_foldrM(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Vector_foldrM___rarg), 4, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Vector_foldrM___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Vector_foldrM(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = lean_apply_2(x_1, x_5, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_3 = x_10; +x_5 = x_8; +goto _start; +} +else +{ +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg___boxed), 5, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Vector_foldl___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_array_get_size(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_4); +if (x_6 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_7; +x_7 = lean_nat_dec_le(x_4, x_4); +if (x_7 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; +} +else +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_10 = l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg(x_1, x_3, x_8, x_9, x_2); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Vector_foldl(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Vector_foldl___rarg___boxed), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Vector_foldl___spec__1___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Vector_foldl___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Vector_foldl___rarg(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Vector_foldl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Vector_foldl(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; +x_7 = 1; +x_8 = lean_usize_sub(x_3, x_7); +x_9 = lean_array_uget(x_2, x_8); +lean_inc(x_1); +x_10 = lean_apply_2(x_1, x_9, x_5); +x_3 = x_8; +x_5 = x_10; +goto _start; +} +else +{ +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg___boxed), 5, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; +x_7 = 1; +x_8 = lean_usize_sub(x_3, x_7); +x_9 = lean_array_uget(x_2, x_8); +lean_inc(x_1); +x_10 = lean_apply_2(x_1, x_9, x_5); +x_3 = x_8; +x_5 = x_10; +goto _start; +} +else +{ +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg___boxed), 5, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Vector_foldr___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_3); +x_5 = lean_nat_dec_le(x_4, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_4); +if (x_7 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; +} +else +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_9 = 0; +x_10 = l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg(x_1, x_3, x_8, x_9, x_2); +return x_10; +} +} +else +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_lt(x_11, x_4); +if (x_12 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; +} +else +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_14 = 0; +x_15 = l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg(x_1, x_3, x_13, x_14, x_2); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Vector_foldr(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Vector_foldr___rarg___boxed), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__1___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldrMUnsafe_fold___at_Vector_foldr___spec__2___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Vector_foldr___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Vector_foldr___rarg(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Vector_foldr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Vector_foldr(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} LEAN_EXPORT lean_object* l_Vector_append___rarg(lean_object* x_1, lean_object* x_2) { _start: { @@ -2944,7 +3363,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1435_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1591_() { _start: { lean_object* x_1; @@ -2952,7 +3371,7 @@ x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__17; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1443_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1599_() { _start: { lean_object* x_1; @@ -3030,7 +3449,7 @@ lean_dec(x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1517_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1673_() { _start: { lean_object* x_1; @@ -3450,7 +3869,7 @@ lean_dec(x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1773_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1929_() { _start: { lean_object* x_1; @@ -3530,7 +3949,7 @@ static lean_object* _init_l_Vector_eraseIdx_x21___rarg___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Vector_eraseIdx_x21___rarg___closed__1; x_2 = l_Vector_eraseIdx_x21___rarg___closed__2; -x_3 = lean_unsigned_to_nat(249u); +x_3 = lean_unsigned_to_nat(261u); x_4 = lean_unsigned_to_nat(4u); x_5 = l_Vector_eraseIdx_x21___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4256,7 +4675,7 @@ lean_dec(x_2); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1() { _start: { lean_object* x_1; @@ -4264,41 +4683,41 @@ x_1 = lean_mk_string_unchecked("exact", 5, 5); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__1; x_2 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__2; x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__1; -x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1; +x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5() { _start: { lean_object* x_1; @@ -4306,19 +4725,19 @@ x_1 = lean_mk_string_unchecked("paren", 5, 5); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__1; x_2 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__2; x_3 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__3; -x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5; +x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4330,17 +4749,17 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9() { _start: { lean_object* x_1; @@ -4348,17 +4767,17 @@ x_1 = lean_mk_string_unchecked("term_<_", 7, 7); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11() { _start: { lean_object* x_1; @@ -4366,19 +4785,19 @@ x_1 = lean_mk_string_unchecked("cdot", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__1; x_2 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__2; x_3 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__3; -x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11; +x_4 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13() { _start: { lean_object* x_1; @@ -4386,35 +4805,35 @@ x_1 = lean_mk_string_unchecked("·", 2, 1); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4422,17 +4841,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18() { _start: { lean_object* x_1; @@ -4440,45 +4859,45 @@ x_1 = lean_mk_string_unchecked("<", 1, 1); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4486,17 +4905,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4508,23 +4927,23 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4532,23 +4951,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4556,23 +4975,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__15; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4580,23 +4999,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__5; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4604,23 +5023,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Vector___aux__Init__Data__Vector__Basic______macroRules__Vector__term_x23v_x5b___x2c_x5d__1___closed__26; -x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32; +x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l___auto____x40_Init_Data_Vector_Basic___hyg_1091____closed__3; -x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33; +x_3 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4628,11 +5047,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196_() { +static lean_object* _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34; +x_1 = l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34; return x_1; } } @@ -5116,12 +5535,12 @@ l_Vector_instHAppendHAddNat___closed__1 = _init_l_Vector_instHAppendHAddNat___cl lean_mark_persistent(l_Vector_instHAppendHAddNat___closed__1); l_Vector_zipWith___rarg___closed__1 = _init_l_Vector_zipWith___rarg___closed__1(); lean_mark_persistent(l_Vector_zipWith___rarg___closed__1); -l___auto____x40_Init_Data_Vector_Basic___hyg_1435_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1435_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1435_); -l___auto____x40_Init_Data_Vector_Basic___hyg_1443_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1443_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1443_); -l___auto____x40_Init_Data_Vector_Basic___hyg_1517_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1517_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1517_); +l___auto____x40_Init_Data_Vector_Basic___hyg_1591_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1591_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1591_); +l___auto____x40_Init_Data_Vector_Basic___hyg_1599_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1599_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1599_); +l___auto____x40_Init_Data_Vector_Basic___hyg_1673_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1673_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1673_); l_Vector_swapAt_x21___rarg___closed__1 = _init_l_Vector_swapAt_x21___rarg___closed__1(); lean_mark_persistent(l_Vector_swapAt_x21___rarg___closed__1); l_Vector_swapAt_x21___rarg___closed__2 = _init_l_Vector_swapAt_x21___rarg___closed__2(); @@ -5132,8 +5551,8 @@ l_Vector_swapAt_x21___rarg___closed__4 = _init_l_Vector_swapAt_x21___rarg___clos lean_mark_persistent(l_Vector_swapAt_x21___rarg___closed__4); l_Vector_range___closed__1 = _init_l_Vector_range___closed__1(); lean_mark_persistent(l_Vector_range___closed__1); -l___auto____x40_Init_Data_Vector_Basic___hyg_1773_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1773_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1773_); +l___auto____x40_Init_Data_Vector_Basic___hyg_1929_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_1929_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_1929_); l_Vector_eraseIdx_x21___rarg___closed__1 = _init_l_Vector_eraseIdx_x21___rarg___closed__1(); lean_mark_persistent(l_Vector_eraseIdx_x21___rarg___closed__1); l_Vector_eraseIdx_x21___rarg___closed__2 = _init_l_Vector_eraseIdx_x21___rarg___closed__2(); @@ -5142,76 +5561,76 @@ l_Vector_eraseIdx_x21___rarg___closed__3 = _init_l_Vector_eraseIdx_x21___rarg___ lean_mark_persistent(l_Vector_eraseIdx_x21___rarg___closed__3); l_Vector_eraseIdx_x21___rarg___closed__4 = _init_l_Vector_eraseIdx_x21___rarg___closed__4(); lean_mark_persistent(l_Vector_eraseIdx_x21___rarg___closed__4); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__1); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__2); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__3); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__4); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__5); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__6); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__7); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__8); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__9); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__10); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__11); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__12); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__13); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__14); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__15); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__16); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__17); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__18); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__19); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__20); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__21); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__22); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__23); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__24); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__25); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__26); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__27); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__28); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__29); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__30); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__31); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__32); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__33); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196____closed__34); -l___auto____x40_Init_Data_Vector_Basic___hyg_2196_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2196_(); -lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2196_); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__1); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__2); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__3); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__4); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__5); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__6); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__7); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__8); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__9); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__10); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__11); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__12); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__13); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__14); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__15); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__16); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__17); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__18); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__19); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__20); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__21); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__22); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__23); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__24); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__25); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__26); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__27); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__28); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__29); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__30); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__31); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__32); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__33); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34 = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352____closed__34); +l___auto____x40_Init_Data_Vector_Basic___hyg_2352_ = _init_l___auto____x40_Init_Data_Vector_Basic___hyg_2352_(); +lean_mark_persistent(l___auto____x40_Init_Data_Vector_Basic___hyg_2352_); l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__1(); lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__1); l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__2 = _init_l_Std_Range_forIn_x27_loop___at_Vector_lex___spec__1___rarg___closed__2(); diff --git a/stage0/stdlib/Init/Grind/Norm.c b/stage0/stdlib/Init/Grind/Norm.c index 3d8f931a3d9a..dce02ea78b0d 100644 --- a/stage0/stdlib/Init/Grind/Norm.c +++ b/stage0/stdlib/Init/Grind/Norm.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Grind.Norm -// Imports: Init.SimpLemmas Init.Classical Init.ByCases +// Imports: Init.SimpLemmas Init.PropLemmas Init.Classical Init.ByCases #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* initialize_Init_SimpLemmas(uint8_t builtin, lean_object*); +lean_object* initialize_Init_PropLemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Classical(uint8_t builtin, lean_object*); lean_object* initialize_Init_ByCases(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -24,6 +25,9 @@ _G_initialized = true; res = initialize_Init_SimpLemmas(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_PropLemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Init_Classical(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Init/Grind/Tactics.c b/stage0/stdlib/Init/Grind/Tactics.c index 4e46c7564577..e31928a9f4f1 100644 --- a/stage0/stdlib/Init/Grind/Tactics.c +++ b/stage0/stdlib/Init/Grind/Tactics.c @@ -13,100 +13,702 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__12; +static lean_object* l_Lean_Parser_Tactic_grind___closed__11; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__4; static lean_object* l_Lean_Parser_Tactic_grind___closed__7; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindBwd; +static lean_object* l_Lean_Parser_Attr_grind___closed__13; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindFwd; extern lean_object* l_Lean_Parser_Tactic_optConfig; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__8; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__4; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__3; +static lean_object* l_Lean_Parser_Attr_grind___closed__4; static lean_object* l_Lean_Grind_instInhabitedConfig___closed__1; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__4; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__4; static lean_object* l_Lean_Parser_Tactic_grind___closed__10; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindEqRhs; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__6; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__8; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__9; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grind; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__5; +LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_125____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__2; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__1; static lean_object* l_Lean_Parser_Tactic_grind___closed__6; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_grind; static lean_object* l_Lean_Grind_instBEqConfig___closed__1; static lean_object* l_Lean_Parser_Tactic_grind___closed__2; LEAN_EXPORT lean_object* l_Lean_Grind_instBEqConfig; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__7; static lean_object* l_Lean_Parser_Tactic_grind___closed__9; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__3; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__5; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__3; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__1; static lean_object* l_Lean_Parser_Tactic_grind___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindEqBoth; +static lean_object* l_Lean_Parser_Attr_grind___closed__10; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Grind_instInhabitedConfig; +static lean_object* l_Lean_Parser_Attr_grind___closed__3; static lean_object* l_Lean_Parser_Tactic_grind___closed__5; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__1; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__5; +static lean_object* l_Lean_Parser_Attr_grind___closed__2; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__7; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__5; static lean_object* l_Lean_Parser_Tactic_grind___closed__3; +static lean_object* l_Lean_Parser_Attr_grind___closed__6; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Attr_grind___closed__11; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__2; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__3; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__3; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__2; +static lean_object* l_Lean_Parser_Attr_grindEq___closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Attr_grindEq; +LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_125_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_grind___closed__4; -LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__11; +static lean_object* l_Lean_Parser_Attr_grind___closed__5; static lean_object* l_Lean_Parser_Tactic_grind___closed__8; +static lean_object* l_Lean_Parser_Attr_grindEqBoth___closed__10; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(lean_object*, lean_object*); -static lean_object* _init_l_Lean_Grind_instInhabitedConfig___closed__1() { +static lean_object* l_Lean_Parser_Attr_grind___closed__12; +static lean_object* l_Lean_Parser_Attr_grind___closed__8; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__4; +static lean_object* l_Lean_Parser_Attr_grind___closed__7; +static lean_object* l_Lean_Parser_Attr_grind___closed__14; +static lean_object* l_Lean_Parser_Attr_grindBwd___closed__5; +static lean_object* l_Lean_Parser_Attr_grindFwd___closed__2; +static lean_object* l_Lean_Parser_Attr_grindEqRhs___closed__2; +static lean_object* l_Lean_Parser_Attr_grind___closed__1; +static lean_object* l_Lean_Parser_Attr_grind___closed__9; +static lean_object* l_Lean_Parser_Tactic_grind___closed__12; +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__1() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 0; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_alloc_ctor(0, 4, 1); -lean_ctor_set(x_3, 0, x_2); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Attr", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEq", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindEq___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("=", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_grindEq___closed__6; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEq___closed__4; +x_2 = l_Lean_Parser_Attr_grindEq___closed__5; +x_3 = l_Lean_Parser_Attr_grindEq___closed__7; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindEq___closed__8; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEqBoth", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindEqBoth___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("atomic", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__7; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__8; +x_3 = l_Lean_Parser_Attr_grindEq___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__9; +x_3 = l_Lean_Parser_Attr_grindEqBoth___closed__8; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__4; +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 2, x_2); -lean_ctor_set(x_3, 3, x_2); -lean_ctor_set_uint8(x_3, sizeof(void*)*4, x_1); return x_3; } } -static lean_object* _init_l_Lean_Grind_instInhabitedConfig() { +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__1; +x_2 = l_Lean_Parser_Attr_grindEqBoth___closed__2; +x_3 = l_Lean_Parser_Attr_grindEqBoth___closed__11; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqBoth() { _start: { lean_object* x_1; -x_1 = l_Lean_Grind_instInhabitedConfig___closed__1; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__12; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__1() { _start: { -uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_1, 1); -x_6 = lean_ctor_get(x_1, 2); -x_7 = lean_ctor_get(x_1, 3); -x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*4); -x_9 = lean_ctor_get(x_2, 0); -x_10 = lean_ctor_get(x_2, 1); -x_11 = lean_ctor_get(x_2, 2); -x_12 = lean_ctor_get(x_2, 3); -if (x_3 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEqRhs", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__2() { +_start: { -if (x_8 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindEqRhs___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__3() { +_start: { -uint8_t x_23; -x_23 = 1; -x_13 = x_23; -goto block_22; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grindEq___closed__7; +x_3 = l_Lean_Parser_Attr_grindEqBoth___closed__8; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } -else +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__4() { +_start: { -uint8_t x_24; -x_24 = 0; -x_13 = x_24; -goto block_22; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__4; +x_2 = l_Lean_Parser_Attr_grindEqRhs___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqRhs___closed__1; +x_2 = l_Lean_Parser_Attr_grindEqRhs___closed__2; +x_3 = l_Lean_Parser_Attr_grindEqRhs___closed__4; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindEqRhs() { +_start: { -if (x_8 == 0) +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindEqRhs___closed__5; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__1() { +_start: { -uint8_t x_25; -x_25 = 0; -x_13 = x_25; -goto block_22; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindBwd", 8, 8); +return x_1; } -else +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindBwd___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("←", 3, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__4() { +_start: { -uint8_t x_26; -x_26 = 1; -x_13 = x_26; -goto block_22; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_grindBwd___closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -block_22: +static lean_object* _init_l_Lean_Parser_Attr_grindBwd___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindBwd___closed__1; +x_2 = l_Lean_Parser_Attr_grindBwd___closed__2; +x_3 = l_Lean_Parser_Attr_grindBwd___closed__4; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindBwd() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindBwd___closed__5; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindFwd", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grindFwd___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__3() { +_start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("→", 3, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_grindFwd___closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindFwd___closed__1; +x_2 = l_Lean_Parser_Attr_grindFwd___closed__2; +x_3 = l_Lean_Parser_Attr_grindFwd___closed__4; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grindFwd() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grindFwd___closed__5; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Attr_grindEq___closed__3; +x_4 = l_Lean_Parser_Attr_grind___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__3() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grind___closed__1; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optional", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Attr_grind___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("orelse", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Attr_grind___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__7; +x_2 = l_Lean_Parser_Attr_grindBwd; +x_3 = l_Lean_Parser_Attr_grindFwd; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__7; +x_2 = l_Lean_Parser_Attr_grindEq; +x_3 = l_Lean_Parser_Attr_grind___closed__8; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__7; +x_2 = l_Lean_Parser_Attr_grindEqRhs; +x_3 = l_Lean_Parser_Attr_grind___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__7; +x_2 = l_Lean_Parser_Attr_grindEqBoth; +x_3 = l_Lean_Parser_Attr_grind___closed__10; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grind___closed__5; +x_2 = l_Lean_Parser_Attr_grind___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grind___closed__3; +x_3 = l_Lean_Parser_Attr_grind___closed__12; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grind___closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Parser_Attr_grind___closed__13; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Attr_grind() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Attr_grind___closed__14; +return x_1; +} +} +static lean_object* _init_l_Lean_Grind_instInhabitedConfig___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*4, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Grind_instInhabitedConfig() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Grind_instInhabitedConfig___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_125_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; uint8_t x_13; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get(x_1, 3); +x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +x_8 = lean_ctor_get(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_uint8(x_2, sizeof(void*)*4); +x_13 = lean_nat_dec_eq(x_3, x_8); if (x_13 == 0) { uint8_t x_14; @@ -145,21 +747,36 @@ return x_20; } else { +if (x_7 == 0) +{ +if (x_12 == 0) +{ uint8_t x_21; -x_21 = lean_nat_dec_eq(x_7, x_12); +x_21 = 1; return x_21; } +else +{ +uint8_t x_22; +x_22 = 0; +return x_22; +} +} +else +{ +return x_12; +} } } } } } } -LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_125____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30_(x_1, x_2); +x_3 = l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_125_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -170,7 +787,7 @@ static lean_object* _init_l_Lean_Grind_instBEqConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_30____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Grind_Tactics_0__Lean_Grind_beqConfig____x40_Init_Grind_Tactics___hyg_125____boxed), 2, 0); return x_1; } } @@ -186,51 +803,59 @@ static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); return x_1; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Parser", 6, 6); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Attr_grindEq___closed__1; +x_2 = l_Lean_Parser_Attr_grindEq___closed__2; +x_3 = l_Lean_Parser_Tactic_grind___closed__1; +x_4 = l_Lean_Parser_Attr_grind___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Tactic", 6, 6); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Attr_grind___closed__3; +x_3 = l_Lean_Parser_Tactic_optConfig; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); +x_1 = lean_mk_string_unchecked("on_failure ", 11, 11); return x_1; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Tactic_grind___closed__1; -x_2 = l_Lean_Parser_Tactic_grind___closed__2; -x_3 = l_Lean_Parser_Tactic_grind___closed__3; -x_4 = l_Lean_Parser_Tactic_grind___closed__4; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_grind___closed__4; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("andthen", 7, 7); +x_1 = lean_mk_string_unchecked("term", 4, 4); return x_1; } } @@ -247,12 +872,12 @@ return x_3; static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__8() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_grind___closed__4; -x_2 = 0; -x_3 = lean_alloc_ctor(6, 1, 1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_grind___closed__7; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +lean_ctor_set(x_3, 1, x_2); return x_3; } } @@ -260,9 +885,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_grind___closed__7; -x_2 = l_Lean_Parser_Tactic_grind___closed__8; -x_3 = l_Lean_Parser_Tactic_optConfig; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Tactic_grind___closed__5; +x_3 = l_Lean_Parser_Tactic_grind___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -273,10 +898,36 @@ return x_4; static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__10() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_grind___closed__5; +x_2 = l_Lean_Parser_Tactic_grind___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Attr_grindEqBoth___closed__6; +x_2 = l_Lean_Parser_Tactic_grind___closed__3; +x_3 = l_Lean_Parser_Tactic_grind___closed__10; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_grind___closed__12() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_grind___closed__5; +x_1 = l_Lean_Parser_Tactic_grind___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_grind___closed__9; +x_3 = l_Lean_Parser_Tactic_grind___closed__11; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -288,7 +939,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_grind() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_grind___closed__10; +x_1 = l_Lean_Parser_Tactic_grind___closed__12; return x_1; } } @@ -301,6 +952,116 @@ _G_initialized = true; res = initialize_Init_Tactics(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_Parser_Attr_grindEq___closed__1 = _init_l_Lean_Parser_Attr_grindEq___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__1); +l_Lean_Parser_Attr_grindEq___closed__2 = _init_l_Lean_Parser_Attr_grindEq___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__2); +l_Lean_Parser_Attr_grindEq___closed__3 = _init_l_Lean_Parser_Attr_grindEq___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__3); +l_Lean_Parser_Attr_grindEq___closed__4 = _init_l_Lean_Parser_Attr_grindEq___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__4); +l_Lean_Parser_Attr_grindEq___closed__5 = _init_l_Lean_Parser_Attr_grindEq___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__5); +l_Lean_Parser_Attr_grindEq___closed__6 = _init_l_Lean_Parser_Attr_grindEq___closed__6(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__6); +l_Lean_Parser_Attr_grindEq___closed__7 = _init_l_Lean_Parser_Attr_grindEq___closed__7(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__7); +l_Lean_Parser_Attr_grindEq___closed__8 = _init_l_Lean_Parser_Attr_grindEq___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq___closed__8); +l_Lean_Parser_Attr_grindEq = _init_l_Lean_Parser_Attr_grindEq(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEq); +l_Lean_Parser_Attr_grindEqBoth___closed__1 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__1); +l_Lean_Parser_Attr_grindEqBoth___closed__2 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__2); +l_Lean_Parser_Attr_grindEqBoth___closed__3 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__3); +l_Lean_Parser_Attr_grindEqBoth___closed__4 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__4); +l_Lean_Parser_Attr_grindEqBoth___closed__5 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__5); +l_Lean_Parser_Attr_grindEqBoth___closed__6 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__6(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__6); +l_Lean_Parser_Attr_grindEqBoth___closed__7 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__7(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__7); +l_Lean_Parser_Attr_grindEqBoth___closed__8 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__8); +l_Lean_Parser_Attr_grindEqBoth___closed__9 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__9); +l_Lean_Parser_Attr_grindEqBoth___closed__10 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__10(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__10); +l_Lean_Parser_Attr_grindEqBoth___closed__11 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__11(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__11); +l_Lean_Parser_Attr_grindEqBoth___closed__12 = _init_l_Lean_Parser_Attr_grindEqBoth___closed__12(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth___closed__12); +l_Lean_Parser_Attr_grindEqBoth = _init_l_Lean_Parser_Attr_grindEqBoth(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqBoth); +l_Lean_Parser_Attr_grindEqRhs___closed__1 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__1); +l_Lean_Parser_Attr_grindEqRhs___closed__2 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__2); +l_Lean_Parser_Attr_grindEqRhs___closed__3 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__3); +l_Lean_Parser_Attr_grindEqRhs___closed__4 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__4); +l_Lean_Parser_Attr_grindEqRhs___closed__5 = _init_l_Lean_Parser_Attr_grindEqRhs___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs___closed__5); +l_Lean_Parser_Attr_grindEqRhs = _init_l_Lean_Parser_Attr_grindEqRhs(); +lean_mark_persistent(l_Lean_Parser_Attr_grindEqRhs); +l_Lean_Parser_Attr_grindBwd___closed__1 = _init_l_Lean_Parser_Attr_grindBwd___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__1); +l_Lean_Parser_Attr_grindBwd___closed__2 = _init_l_Lean_Parser_Attr_grindBwd___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__2); +l_Lean_Parser_Attr_grindBwd___closed__3 = _init_l_Lean_Parser_Attr_grindBwd___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__3); +l_Lean_Parser_Attr_grindBwd___closed__4 = _init_l_Lean_Parser_Attr_grindBwd___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__4); +l_Lean_Parser_Attr_grindBwd___closed__5 = _init_l_Lean_Parser_Attr_grindBwd___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd___closed__5); +l_Lean_Parser_Attr_grindBwd = _init_l_Lean_Parser_Attr_grindBwd(); +lean_mark_persistent(l_Lean_Parser_Attr_grindBwd); +l_Lean_Parser_Attr_grindFwd___closed__1 = _init_l_Lean_Parser_Attr_grindFwd___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__1); +l_Lean_Parser_Attr_grindFwd___closed__2 = _init_l_Lean_Parser_Attr_grindFwd___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__2); +l_Lean_Parser_Attr_grindFwd___closed__3 = _init_l_Lean_Parser_Attr_grindFwd___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__3); +l_Lean_Parser_Attr_grindFwd___closed__4 = _init_l_Lean_Parser_Attr_grindFwd___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__4); +l_Lean_Parser_Attr_grindFwd___closed__5 = _init_l_Lean_Parser_Attr_grindFwd___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd___closed__5); +l_Lean_Parser_Attr_grindFwd = _init_l_Lean_Parser_Attr_grindFwd(); +lean_mark_persistent(l_Lean_Parser_Attr_grindFwd); +l_Lean_Parser_Attr_grind___closed__1 = _init_l_Lean_Parser_Attr_grind___closed__1(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__1); +l_Lean_Parser_Attr_grind___closed__2 = _init_l_Lean_Parser_Attr_grind___closed__2(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__2); +l_Lean_Parser_Attr_grind___closed__3 = _init_l_Lean_Parser_Attr_grind___closed__3(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__3); +l_Lean_Parser_Attr_grind___closed__4 = _init_l_Lean_Parser_Attr_grind___closed__4(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__4); +l_Lean_Parser_Attr_grind___closed__5 = _init_l_Lean_Parser_Attr_grind___closed__5(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__5); +l_Lean_Parser_Attr_grind___closed__6 = _init_l_Lean_Parser_Attr_grind___closed__6(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__6); +l_Lean_Parser_Attr_grind___closed__7 = _init_l_Lean_Parser_Attr_grind___closed__7(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__7); +l_Lean_Parser_Attr_grind___closed__8 = _init_l_Lean_Parser_Attr_grind___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__8); +l_Lean_Parser_Attr_grind___closed__9 = _init_l_Lean_Parser_Attr_grind___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__9); +l_Lean_Parser_Attr_grind___closed__10 = _init_l_Lean_Parser_Attr_grind___closed__10(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__10); +l_Lean_Parser_Attr_grind___closed__11 = _init_l_Lean_Parser_Attr_grind___closed__11(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__11); +l_Lean_Parser_Attr_grind___closed__12 = _init_l_Lean_Parser_Attr_grind___closed__12(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__12); +l_Lean_Parser_Attr_grind___closed__13 = _init_l_Lean_Parser_Attr_grind___closed__13(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__13); +l_Lean_Parser_Attr_grind___closed__14 = _init_l_Lean_Parser_Attr_grind___closed__14(); +lean_mark_persistent(l_Lean_Parser_Attr_grind___closed__14); +l_Lean_Parser_Attr_grind = _init_l_Lean_Parser_Attr_grind(); +lean_mark_persistent(l_Lean_Parser_Attr_grind); l_Lean_Grind_instInhabitedConfig___closed__1 = _init_l_Lean_Grind_instInhabitedConfig___closed__1(); lean_mark_persistent(l_Lean_Grind_instInhabitedConfig___closed__1); l_Lean_Grind_instInhabitedConfig = _init_l_Lean_Grind_instInhabitedConfig(); @@ -329,6 +1090,10 @@ l_Lean_Parser_Tactic_grind___closed__9 = _init_l_Lean_Parser_Tactic_grind___clos lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__9); l_Lean_Parser_Tactic_grind___closed__10 = _init_l_Lean_Parser_Tactic_grind___closed__10(); lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__10); +l_Lean_Parser_Tactic_grind___closed__11 = _init_l_Lean_Parser_Tactic_grind___closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__11); +l_Lean_Parser_Tactic_grind___closed__12 = _init_l_Lean_Parser_Tactic_grind___closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic_grind___closed__12); l_Lean_Parser_Tactic_grind = _init_l_Lean_Parser_Tactic_grind(); lean_mark_persistent(l_Lean_Parser_Tactic_grind); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Init/Grind/Util.c b/stage0/stdlib/Init/Grind/Util.c index 02d8bb9beb3d..9a31181dd992 100644 --- a/stage0/stdlib/Init/Grind/Util.c +++ b/stage0/stdlib/Init/Grind/Util.c @@ -13,6 +13,54 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_offset(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_offset___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp___rarg(lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Grind_doNotSimp___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_doNotSimp___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Grind_doNotSimp___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_offset(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Grind_offset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Grind_offset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} lean_object* initialize_Init_Core(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object* w) { diff --git a/stage0/stdlib/Init/Prelude.c b/stage0/stdlib/Init/Prelude.c index 350b5f3c331a..3dd6be75eba9 100644 --- a/stage0/stdlib/Init/Prelude.c +++ b/stage0/stdlib/Init/Prelude.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Lean_withRef_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_instMulNat___closed__1; LEAN_EXPORT lean_object* l_instInhabitedReaderT___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_read(lean_object*, lean_object*); @@ -729,6 +730,7 @@ LEAN_EXPORT lean_object* l_decEq(lean_object*); LEAN_EXPORT lean_object* l_EStateM_instOrElseOfBacktrackable(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadRefOfMonadLiftOfMonadFunctor___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadLiftTOfMonadLift(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withRef_x3f(lean_object*); LEAN_EXPORT lean_object* l_instLTNat; LEAN_EXPORT lean_object* l_instDecidableAnd(lean_object*, lean_object*); static lean_object* l_EStateM_instMonad___closed__6; @@ -9603,6 +9605,43 @@ lean_dec(x_1); return x_5; } } +LEAN_EXPORT lean_object* l_Lean_withRef_x3f___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +if (lean_obj_tag(x_4) == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_dec(x_1); +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l_Lean_withRef___rarg___lambda__1___boxed), 4, 3); +lean_closure_set(x_9, 0, x_6); +lean_closure_set(x_9, 1, x_2); +lean_closure_set(x_9, 2, x_5); +x_10 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_8, x_9); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withRef_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_withRef_x3f___rarg), 5, 0); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { diff --git a/stage0/stdlib/Init/Tactics.c b/stage0/stdlib/Init/Tactics.c index 1aae711684f2..1c0cc09ecdfc 100644 --- a/stage0/stdlib/Init/Tactics.c +++ b/stage0/stdlib/Init/Tactics.c @@ -16374,7 +16374,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_revert___closed__5; +x_1 = l_Lean_Parser_Tactic_intro___closed__11; x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); diff --git a/stage0/stdlib/Lake/CLI/Translate.c b/stage0/stdlib/Lake/CLI/Translate.c index cbc42c813c37..3053bdfcbcd0 100644 --- a/stage0/stdlib/Lake/CLI/Translate.c +++ b/stage0/stdlib/Lake/CLI/Translate.c @@ -35,8 +35,6 @@ static lean_object* l_Lake_Package_mkConfigString___closed__21; extern lean_object* l_Lean_maxRecDepth; lean_object* l_Lake_importModulesUsingCache(lean_object*, lean_object*, uint32_t, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lake_Package_mkConfigString___closed__13; LEAN_EXPORT lean_object* l_Lake_Package_mkConfigString(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -45,6 +43,7 @@ lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(le static lean_object* l_Lake_Package_mkConfigString___closed__12; static lean_object* l_Lake_Package_mkConfigString___closed__14; LEAN_EXPORT lean_object* l_Lake_Package_mkConfigString___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l___private_Lake_CLI_Translate_0__Lake_descopeTSyntax___rarg(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lake_Package_mkConfigString___closed__16; @@ -90,6 +89,7 @@ lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lake_Package_mkConfigString___closed__5; static lean_object* l_Lake_Package_mkConfigString___closed__17; lean_object* lean_array_get_size(lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lake_Package_mkConfigString___closed__19; @@ -780,7 +780,7 @@ lean_dec(x_171); x_174 = lean_ctor_get(x_172, 0); lean_inc(x_174); lean_dec(x_172); -x_175 = l_Lean_Kernel_isDiagnosticsEnabled(x_174); +x_175 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_174); lean_dec(x_174); if (x_175 == 0) { @@ -1008,7 +1008,7 @@ x_181 = lean_ctor_get(x_178, 0); x_182 = lean_ctor_get(x_178, 4); lean_dec(x_182); x_183 = l_Lake_Package_mkConfigString___closed__28; -x_184 = l_Lean_Kernel_enableDiag(x_181, x_183); +x_184 = l_Lean_Kernel_Environment_enableDiag(x_181, x_183); lean_ctor_set(x_178, 4, x_105); lean_ctor_set(x_178, 0, x_184); x_185 = lean_st_ref_set(x_169, x_178, x_179); @@ -1123,7 +1123,7 @@ lean_inc(x_208); lean_inc(x_207); lean_dec(x_178); x_214 = l_Lake_Package_mkConfigString___closed__28; -x_215 = l_Lean_Kernel_enableDiag(x_207, x_214); +x_215 = l_Lean_Kernel_Environment_enableDiag(x_207, x_214); x_216 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_216, 0, x_215); lean_ctor_set(x_216, 1, x_208); diff --git a/stage0/stdlib/Lake/DSL/Config.c b/stage0/stdlib/Lake/DSL/Config.c index 083f1fdf4eea..9cb38717ba8a 100644 --- a/stage0/stdlib/Lake/DSL/Config.c +++ b/stage0/stdlib/Lake/DSL/Config.c @@ -95,7 +95,7 @@ static lean_object* l_Lake_DSL_elabGetConfig___closed__15; static lean_object* l_Lake_DSL_getConfig___closed__11; static lean_object* l_Lake_DSL_dirConst___closed__7; static lean_object* l_Lake_DSL_elabGetConfig___closed__5; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lake_DSL_elabGetConfig___closed__37; static lean_object* l_Lake_DSL_getConfig___closed__2; static lean_object* l_Lake_DSL_elabGetConfig___closed__42; @@ -1209,7 +1209,8 @@ x_67 = lean_ctor_get(x_64, 1); x_68 = lean_ctor_get(x_66, 0); lean_inc(x_68); lean_dec(x_66); -x_69 = lean_environment_main_module(x_68); +x_69 = l_Lean_Environment_mainModule(x_68); +lean_dec(x_68); x_70 = l_Lake_DSL_elabGetConfig___closed__13; lean_inc(x_62); lean_ctor_set_tag(x_64, 2); @@ -1292,7 +1293,8 @@ lean_dec(x_64); x_103 = lean_ctor_get(x_101, 0); lean_inc(x_103); lean_dec(x_101); -x_104 = lean_environment_main_module(x_103); +x_104 = l_Lean_Environment_mainModule(x_103); +lean_dec(x_103); x_105 = l_Lake_DSL_elabGetConfig___closed__13; lean_inc(x_62); x_106 = lean_alloc_ctor(2, 2, 0); @@ -1388,7 +1390,8 @@ lean_dec(x_142); x_145 = lean_ctor_get(x_143, 0); lean_inc(x_145); lean_dec(x_143); -x_146 = lean_environment_main_module(x_145); +x_146 = l_Lean_Environment_mainModule(x_145); +lean_dec(x_145); x_147 = l_Lake_DSL_elabGetConfig___closed__45; x_148 = l_Lean_addMacroScope(x_146, x_147, x_141); x_149 = l_Lake_DSL_elabGetConfig___closed__44; @@ -1546,7 +1549,8 @@ if (lean_is_exclusive(x_209)) { x_213 = lean_ctor_get(x_210, 0); lean_inc(x_213); lean_dec(x_210); -x_214 = lean_environment_main_module(x_213); +x_214 = l_Lean_Environment_mainModule(x_213); +lean_dec(x_213); x_215 = l_Lake_DSL_elabGetConfig___closed__13; lean_inc(x_207); if (lean_is_scalar(x_212)) { @@ -1645,7 +1649,8 @@ lean_dec(x_253); x_256 = lean_ctor_get(x_254, 0); lean_inc(x_256); lean_dec(x_254); -x_257 = lean_environment_main_module(x_256); +x_257 = l_Lean_Environment_mainModule(x_256); +lean_dec(x_256); x_258 = l_Lake_DSL_elabGetConfig___closed__45; x_259 = l_Lean_addMacroScope(x_257, x_258, x_252); x_260 = l_Lake_DSL_elabGetConfig___closed__44; diff --git a/stage0/stdlib/Lake/Load/Lean/Elab.c b/stage0/stdlib/Lake/Load/Lean/Elab.c index dc13c42845c8..d6dd92eb6611 100644 --- a/stage0/stdlib/Lake/Load/Lean/Elab.c +++ b/stage0/stdlib/Lake/Load/Lean/Elab.c @@ -134,7 +134,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_System_FilePath_fileName(lean_object*); static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__34; static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__23; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static uint64_t l_Lake_importModulesUsingCache___closed__3; static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__18; static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__5; @@ -143,7 +142,7 @@ static lean_object* l_Lake_processHeader___closed__1; static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__38; LEAN_EXPORT uint64_t l___private_Lake_Load_Lean_Elab_0__Lake_hashImport____x40_Lake_Load_Lean_Elab___hyg_80_(lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lake_importModulesUsingCache___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_set_main_module(lean_object*, lean_object*); +lean_object* l_Lean_Environment_setMainModule(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lake_Load_Lean_Elab_0__Lake_fromJsonConfigTrace____x40_Lake_Load_Lean_Elab___hyg_1110____spec__1(lean_object*, lean_object*); static lean_object* l_Lake_instFromJsonConfigTrace___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lake_importConfigFileCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -282,9 +281,10 @@ LEAN_EXPORT lean_object* l_Lake_importModulesUsingCache___boxed(lean_object*, le static lean_object* l___private_Lake_Load_Lean_Elab_0__Lake_fromJsonConfigTrace____x40_Lake_Load_Lean_Elab___hyg_1110____closed__15; static lean_object* l___private_Lake_Load_Lean_Elab_0__Lake_fromJsonConfigTrace____x40_Lake_Load_Lean_Elab___hyg_1110____closed__21; static lean_object* l___private_Lake_Load_Lean_Elab_0__Lake_fromJsonConfigTrace____x40_Lake_Load_Lean_Elab___hyg_1110____closed__14; -lean_object* lean_environment_add(lean_object*, lean_object*); +lean_object* lean_elab_environment_add(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lake_importConfigFileCore_lakeExts___closed__28; +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lake_Load_Lean_Elab_0__Lake_addToEnv___boxed(lean_object*, lean_object*); lean_object* l_Lean_MessageLog_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lake_importModulesUsingCache___spec__3(lean_object*); @@ -2873,7 +2873,7 @@ x_101 = lean_ctor_get(x_99, 1); lean_inc(x_101); lean_dec(x_99); x_102 = l_Lake_configModuleName; -x_103 = lean_environment_set_main_module(x_100, x_102); +x_103 = l_Lean_Environment_setMainModule(x_100, x_102); x_104 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_104, 0, x_1); x_105 = l_Lake_elabConfigFile___closed__2; @@ -2908,7 +2908,7 @@ x_117 = lean_ctor_get(x_114, 1); lean_inc(x_117); lean_dec(x_114); x_118 = l_Lake_configModuleName; -x_119 = lean_environment_set_main_module(x_116, x_118); +x_119 = l_Lean_Environment_setMainModule(x_116, x_118); x_120 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_120, 0, x_1); x_121 = l_Lake_elabConfigFile___closed__2; @@ -3063,7 +3063,7 @@ x_159 = lean_ctor_get(x_155, 1); lean_inc(x_159); lean_dec(x_155); x_160 = l_Lake_configModuleName; -x_161 = lean_environment_set_main_module(x_158, x_160); +x_161 = l_Lean_Environment_setMainModule(x_158, x_160); x_162 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_162, 0, x_1); x_163 = l_Lake_elabConfigFile___closed__2; @@ -3324,7 +3324,7 @@ x_228 = lean_ctor_get(x_224, 1); lean_inc(x_228); lean_dec(x_224); x_229 = l_Lake_configModuleName; -x_230 = lean_environment_set_main_module(x_227, x_229); +x_230 = l_Lean_Environment_setMainModule(x_227, x_229); x_231 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_231, 0, x_1); x_232 = l_Lake_elabConfigFile___closed__2; @@ -3476,7 +3476,7 @@ LEAN_EXPORT lean_object* l___private_Lake_Load_Lean_Elab_0__Lake_addToEnv___boxe _start: { lean_object* x_3; -x_3 = lean_environment_add(x_1, x_2); +x_3 = lean_elab_environment_add(x_1, x_2); return x_3; } } @@ -4221,7 +4221,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; x_6 = lean_array_uget(x_1, x_2); -x_7 = lean_environment_add(x_4, x_6); +x_7 = lean_elab_environment_add(x_4, x_6); x_8 = 1; x_9 = lean_usize_add(x_2, x_8); x_2 = x_9; @@ -4643,7 +4643,7 @@ x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_10); lean_ctor_set(x_30, 1, x_29); x_31 = l___private_Lake_Load_Lean_Elab_0__Lake_toJsonConfigTrace____x40_Lake_Load_Lean_Elab___hyg_1030____closed__4; -x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_30, x_31); +x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_30, x_31); x_33 = l_Lean_Json_mkObj(x_32); return x_33; } diff --git a/stage0/stdlib/Lake/Load/Lean/Eval.c b/stage0/stdlib/Lake/Load/Lean/Eval.c index 443b1dc53ac7..bed6a6cba4c0 100644 --- a/stage0/stdlib/Lake/Load/Lean/Eval.c +++ b/stage0/stdlib/Lake/Load/Lean/Eval.c @@ -60,7 +60,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lake_Package_loadFromEnv___spec_ lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lake_PackageConfig_loadFromEnv___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lake_Package_loadFromEnv___lambda__1___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lake_Workspace_addFacetsFromEnv___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lake_unsafeEvalConstCheck_throwUnexpectedType(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lake_Package_loadFromEnv___spec__21___closed__1; @@ -338,9 +338,8 @@ LEAN_EXPORT lean_object* l_Lake_unsafeEvalConstCheck(lean_object* x_1, lean_obje _start: { lean_object* x_6; -lean_inc(x_5); lean_inc(x_1); -x_6 = lean_environment_find(x_1, x_5); +x_6 = l_Lean_Environment_find_x3f(x_1, x_5); if (lean_obj_tag(x_6) == 0) { uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; diff --git a/stage0/stdlib/Lake/Toml/Load.c b/stage0/stdlib/Lake/Toml/Load.c index 2b8cd788ea58..11ce061f0914 100644 --- a/stage0/stdlib/Lake/Toml/Load.c +++ b/stage0/stdlib/Lake/Toml/Load.c @@ -31,8 +31,6 @@ extern lean_object* l_Lean_maxRecDepth; extern lean_object* l_Lake_Toml_toml; LEAN_EXPORT lean_object* l_Lake_Toml_loadToml___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserFn_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lake_Toml_loadToml___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lake_Toml_loadToml___lambda__2___closed__10; @@ -40,6 +38,7 @@ lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(le static lean_object* l_Lake_Toml_loadToml___lambda__2___closed__5; static lean_object* l_Lake_Toml_loadToml___lambda__2___closed__16; lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); lean_object* l_Lean_Parser_mkParserState(lean_object*); static lean_object* l_Lake_Toml_loadToml___lambda__2___closed__2; static lean_object* l_Lake_Toml_loadToml___lambda__2___closed__3; @@ -70,6 +69,7 @@ extern lean_object* l_Lean_firstFrontendMacroScope; lean_object* lean_array_mk(lean_object*); lean_object* lean_io_error_to_string(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); lean_object* lean_mk_empty_environment(uint32_t, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_Lake_mkMessageNoPos(lean_object*, lean_object*, uint8_t); @@ -493,7 +493,7 @@ lean_dec(x_64); x_67 = lean_ctor_get(x_65, 0); lean_inc(x_67); lean_dec(x_65); -x_68 = l_Lean_Kernel_isDiagnosticsEnabled(x_67); +x_68 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_67); lean_dec(x_67); if (x_68 == 0) { @@ -548,7 +548,7 @@ x_74 = lean_ctor_get(x_71, 0); x_75 = lean_ctor_get(x_71, 4); lean_dec(x_75); x_76 = l_Lake_Toml_loadToml___lambda__2___closed__22; -x_77 = l_Lean_Kernel_enableDiag(x_74, x_76); +x_77 = l_Lean_Kernel_Environment_enableDiag(x_74, x_76); lean_ctor_set(x_71, 4, x_56); lean_ctor_set(x_71, 0, x_77); x_78 = lean_st_ref_set(x_62, x_71, x_72); @@ -635,7 +635,7 @@ lean_inc(x_97); lean_inc(x_96); lean_dec(x_71); x_103 = l_Lake_Toml_loadToml___lambda__2___closed__22; -x_104 = l_Lean_Kernel_enableDiag(x_96, x_103); +x_104 = l_Lean_Kernel_Environment_enableDiag(x_96, x_103); x_105 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_105, 0, x_104); lean_ctor_set(x_105, 1, x_97); diff --git a/stage0/stdlib/Lake/Util/Version.c b/stage0/stdlib/Lake/Util/Version.c index a31ce9b2239e..ab863cb24aac 100644 --- a/stage0/stdlib/Lake/Util/Version.c +++ b/stage0/stdlib/Lake/Util/Version.c @@ -248,7 +248,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lake_elabVerLit___closed__7; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lake_instToExprSemVerCore___lambda__1(lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Substring_takeRightWhileAux___at_Substring_trimRight___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lake_Util_Version_0__Lake_reprSemVerCore____x40_Lake_Util_Version___hyg_48____closed__3; static lean_object* l___private_Lake_Util_Version_0__Lake_reprToolchainVer____x40_Lake_Util_Version___hyg_1626____closed__5; @@ -5344,7 +5344,8 @@ x_38 = lean_ctor_get(x_35, 1); x_39 = lean_ctor_get(x_37, 0); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); +x_40 = l_Lean_Environment_mainModule(x_39); +lean_dec(x_39); x_41 = l_Lake_elabVerLit___closed__15; x_42 = l_Lean_addMacroScope(x_40, x_41, x_34); x_43 = l_Lake_elabVerLit___closed__14; @@ -5633,7 +5634,8 @@ lean_dec(x_35); x_103 = lean_ctor_get(x_101, 0); lean_inc(x_103); lean_dec(x_101); -x_104 = lean_environment_main_module(x_103); +x_104 = l_Lean_Environment_mainModule(x_103); +lean_dec(x_103); x_105 = l_Lake_elabVerLit___closed__15; x_106 = l_Lean_addMacroScope(x_104, x_105, x_34); x_107 = l_Lake_elabVerLit___closed__14; @@ -5991,7 +5993,8 @@ if (lean_is_exclusive(x_180)) { x_184 = lean_ctor_get(x_181, 0); lean_inc(x_184); lean_dec(x_181); -x_185 = lean_environment_main_module(x_184); +x_185 = l_Lean_Environment_mainModule(x_184); +lean_dec(x_184); x_186 = l_Lake_elabVerLit___closed__15; x_187 = l_Lean_addMacroScope(x_185, x_186, x_179); x_188 = l_Lake_elabVerLit___closed__14; diff --git a/stage0/stdlib/Lean/AddDecl.c b/stage0/stdlib/Lean/AddDecl.c index 894099034125..246a8d989381 100644 --- a/stage0/stdlib/Lean/AddDecl.c +++ b/stage0/stdlib/Lean/AddDecl.c @@ -20,10 +20,11 @@ LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__2(lean_object*, lean_object*, static double l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__2___closed__1; lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_add_decl(lean_object*, size_t, lean_object*, lean_object*); +lean_object* lean_elab_add_decl(lean_object*, size_t, lean_object*, lean_object*); lean_object* l_Lean_log___at_Lean_Core_wrapAsyncAsSnapshot___spec__13(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double lean_float_div(double, double); +lean_object* lean_add_decl(lean_object*, size_t, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__4___closed__3; lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Core_wrapAsyncAsSnapshot___spec__3___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -71,27 +72,30 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2___boxed(l LEAN_EXPORT lean_object* l_Lean_Environment_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_AddDecl___hyg_5____closed__6; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); -static lean_object* l_Lean_Environment_addDecl___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___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_Environment_addAndCompile(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__3___closed__2; static lean_object* l_Lean_logWarning___at_Lean_addDecl___spec__3___closed__1; static lean_object* l_Lean_initFn____x40_Lean_AddDecl___hyg_5____closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_addDecl___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_add_decl_without_checking(lean_object*, lean_object*); +lean_object* lean_elab_add_decl_without_checking(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_AddDecl___hyg_5____closed__5; static lean_object* l_Lean_initFn____x40_Lean_AddDecl___hyg_5____closed__1; static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__3(lean_object*, uint8_t, 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_addDecl___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_add_decl_without_checking(lean_object*, lean_object*); uint8_t l_Lean_Declaration_foldExprM___at_Lean_Declaration_hasSorry___spec__1(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___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_Lean_Kernel_Environment_addDecl___closed__1; LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_addDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__6(lean_object*); extern lean_object* l_Lean_trace_profiler; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -104,9 +108,9 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lamb LEAN_EXPORT lean_object* l_Lean_addAndCompile(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_AddDecl___hyg_5____closed__3; -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -192,7 +196,7 @@ x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_5__ return x_5; } } -static lean_object* _init_l_Lean_Environment_addDecl___closed__1() { +static lean_object* _init_l_Lean_Kernel_Environment_addDecl___closed__1() { _start: { lean_object* x_1; @@ -200,11 +204,11 @@ x_1 = l_Lean_debug_skipKernelTC; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Environment_addDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_addDecl(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_Environment_addDecl___closed__1; +x_5 = l_Lean_Kernel_Environment_addDecl___closed__1; x_6 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_2, x_5); if (x_6 == 0) { @@ -223,6 +227,40 @@ return x_10; } } } +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_addDecl___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_Kernel_Environment_addDecl(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Environment_addDecl(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_Kernel_Environment_addDecl___closed__1; +x_6 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_2, x_5); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; lean_object* x_9; +x_7 = l_Lean_Core_getMaxHeartbeats(x_2); +x_8 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_9 = lean_elab_add_decl(x_1, x_8, x_3, x_4); +return x_9; +} +else +{ +lean_object* x_10; +x_10 = lean_elab_add_decl_without_checking(x_1, x_3); +return x_10; +} +} +} LEAN_EXPORT lean_object* l_Lean_Environment_addDecl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -324,7 +362,7 @@ LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1 lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_2, 2); lean_inc(x_5); -x_6 = l_Lean_KernelException_toMessageData(x_1, x_5); +x_6 = l_Lean_Kernel_Exception_toMessageData(x_1, x_5); x_7 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_6, x_2, x_3, x_4); lean_dec(x_2); return x_7; @@ -2265,8 +2303,8 @@ if (lean_io_result_is_error(res)) return res; l_Lean_debug_skipKernelTC = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_debug_skipKernelTC); lean_dec_ref(res); -}l_Lean_Environment_addDecl___closed__1 = _init_l_Lean_Environment_addDecl___closed__1(); -lean_mark_persistent(l_Lean_Environment_addDecl___closed__1); +}l_Lean_Kernel_Environment_addDecl___closed__1 = _init_l_Lean_Kernel_Environment_addDecl___closed__1(); +lean_mark_persistent(l_Lean_Kernel_Environment_addDecl___closed__1); l_Lean_logWarning___at_Lean_addDecl___spec__3___closed__1 = _init_l_Lean_logWarning___at_Lean_addDecl___spec__3___closed__1(); lean_mark_persistent(l_Lean_logWarning___at_Lean_addDecl___spec__3___closed__1); l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__2___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__4___lambda__2___closed__1(); diff --git a/stage0/stdlib/Lean/Attributes.c b/stage0/stdlib/Lean/Attributes.c index d08486d9b0c9..d4ce2b6d5600 100644 --- a/stage0/stdlib/Lean/Attributes.c +++ b/stage0/stdlib/Lean/Attributes.c @@ -91,7 +91,7 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_upda static lean_object* l_Lean_isAttribute___closed__1; LEAN_EXPORT lean_object* l_Lean_EnumAttributes_getValue(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3785____closed__4; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, 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_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_mkAttributeImplOfEntry___spec__1___boxed(lean_object*, lean_object*); @@ -9015,9 +9015,8 @@ LEAN_EXPORT lean_object* l_Lean_mkAttributeImplOfConstantUnsafe(lean_object* x_1 _start: { lean_object* x_4; -lean_inc(x_3); lean_inc(x_1); -x_4 = lean_environment_find(x_1, x_3); +x_4 = l_Lean_Environment_find_x3f(x_1, x_3); if (lean_obj_tag(x_4) == 0) { uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; diff --git a/stage0/stdlib/Lean/BuiltinDocAttr.c b/stage0/stdlib/Lean/BuiltinDocAttr.c index c479d0196197..2e2bc4bdf56b 100644 --- a/stage0/stdlib/Lean/BuiltinDocAttr.c +++ b/stage0/stdlib/Lean/BuiltinDocAttr.c @@ -263,7 +263,6 @@ lean_dec(x_5); x_8 = lean_ctor_get(x_6, 0); lean_inc(x_8); lean_dec(x_6); -lean_inc(x_1); x_9 = l_Lean_isRec___at_Lean_declareBuiltinDocStringAndRanges___spec__2(x_1, x_2, x_3, x_7); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); @@ -1018,6 +1017,7 @@ lean_object* x_5; x_5 = l_Lean_isRec___at_Lean_declareBuiltinDocStringAndRanges___spec__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Class.c b/stage0/stdlib/Lean/Class.c index 20ff814c03c7..e60d080351bc 100644 --- a/stage0/stdlib/Lean/Class.c +++ b/stage0/stdlib/Lean/Class.c @@ -54,7 +54,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Class_0__Lea LEAN_EXPORT uint8_t l_Lean_Expr_hasAnyFVar_visit___at___private_Lean_Class_0__Lean_checkOutParam___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_mkOutParamArgsImplicit___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Class___hyg_83____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Class___hyg_789____closed__6; static lean_object* l___private_Lean_Class_0__Lean_checkOutParam___closed__1; @@ -3539,9 +3539,8 @@ LEAN_EXPORT lean_object* l_Lean_addClass___lambda__2(lean_object* x_1, lean_obje _start: { lean_object* x_4; lean_object* x_13; -lean_inc(x_2); lean_inc(x_1); -x_13 = lean_environment_find(x_1, x_2); +x_13 = l_Lean_Environment_find_x3f(x_1, x_2); 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; diff --git a/stage0/stdlib/Lean/Compiler/CSimpAttr.c b/stage0/stdlib/Lean/Compiler/CSimpAttr.c index 3e7264c10333..276c423c0ca8 100644 --- a/stage0/stdlib/Lean/Compiler/CSimpAttr.c +++ b/stage0/stdlib/Lean/Compiler/CSimpAttr.c @@ -47,7 +47,7 @@ lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_instInhabitedState; static lean_object* l_Lean_Compiler_CSimp_instInhabitedState___closed__5; LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_137_(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_471____closed__19; static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_471____closed__16; @@ -99,6 +99,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_CSimpAttr_0__Lean_Compiler_CS LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_471____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_471____closed__20; lean_object* l_Lean_Expr_appFn_x21(lean_object*); +lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__4(lean_object*); static lean_object* l_Lean_Compiler_CSimp_instInhabitedState___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_137____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_137____spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -119,7 +120,6 @@ static lean_object* l_Lean_Compiler_CSimp_instInhabitedState___closed__4; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_CSimpAttr_0__Lean_Compiler_CSimp_isConstantReplacement_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_add___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__4(lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_137____closed__6; lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); @@ -330,7 +330,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_6600____spec__4(x_4); +x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__4(x_4); lean_ctor_set(x_1, 1, x_6); lean_ctor_set(x_1, 0, x_5); return x_1; @@ -344,7 +344,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_6600____spec__4(x_8); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7023____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); @@ -1680,8 +1680,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -1721,8 +1720,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Compiler/ExternAttr.c b/stage0/stdlib/Lean/Compiler/ExternAttr.c index c6f6902f3e82..8d97af73f496 100644 --- a/stage0/stdlib/Lean/Compiler/ExternAttr.c +++ b/stage0/stdlib/Lean/Compiler/ExternAttr.c @@ -47,7 +47,7 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Compiler_Extern LEAN_EXPORT lean_object* l_Lean_getExternEntryForAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Compiler_ExternAttr_0__Lean_beqExternAttrData____x40_Lean_Compiler_ExternAttr___hyg_382____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqExternEntry; lean_object* l_List_getD___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_List_foldl___at___private_Lean_Compiler_ExternAttr_0__Lean_hashExternAttrData____x40_Lean_Compiler_ExternAttr___hyg_456____spec__1(uint64_t, lean_object*); @@ -67,9 +67,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternC lean_object* lean_string_utf8_byte_size(lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_1152____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_getExternConstArityExport___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); extern lean_object* l_Lean_projectionFnInfoExt; LEAN_EXPORT lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_1152____closed__9; @@ -97,6 +95,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Compil static lean_object* l_Lean_getExternConstArityExport___closed__8; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__12; LEAN_EXPORT uint8_t l_Lean_isExternC(lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___spec__3___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); @@ -195,6 +194,7 @@ static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_1152____ lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__7; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_externAttr; LEAN_EXPORT lean_object* l_Lean_isExtern___boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); @@ -1385,7 +1385,6 @@ x_13 = l_Lean_MapDeclarationExtension_contains___rarg(x_11, x_12, x_10, x_1); if (x_13 == 0) { uint8_t x_14; -lean_inc(x_1); lean_inc(x_10); x_14 = l_Lean_Environment_isConstructor(x_10, x_1); if (x_14 == 0) @@ -1400,9 +1399,8 @@ return x_6; else { lean_object* x_16; -lean_inc(x_1); lean_inc(x_10); -x_16 = lean_environment_find(x_10, x_1); +x_16 = l_Lean_Environment_find_x3f(x_10, x_1); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; @@ -1442,9 +1440,8 @@ return x_22; else { lean_object* x_23; -lean_inc(x_1); lean_inc(x_10); -x_23 = lean_environment_find(x_10, x_1); +x_23 = l_Lean_Environment_find_x3f(x_10, x_1); if (lean_obj_tag(x_23) == 0) { lean_object* x_24; lean_object* x_25; @@ -1499,7 +1496,6 @@ x_35 = l_Lean_MapDeclarationExtension_contains___rarg(x_33, x_34, x_32, x_1); if (x_35 == 0) { uint8_t x_36; -lean_inc(x_1); lean_inc(x_32); x_36 = l_Lean_Environment_isConstructor(x_32, x_1); if (x_36 == 0) @@ -1516,9 +1512,8 @@ return x_38; else { lean_object* x_39; -lean_inc(x_1); lean_inc(x_32); -x_39 = lean_environment_find(x_32, x_1); +x_39 = l_Lean_Environment_find_x3f(x_32, x_1); if (lean_obj_tag(x_39) == 0) { lean_object* x_40; lean_object* x_41; @@ -1558,9 +1553,8 @@ return x_46; else { lean_object* x_47; -lean_inc(x_1); lean_inc(x_32); -x_47 = lean_environment_find(x_32, x_1); +x_47 = l_Lean_Environment_find_x3f(x_32, x_1); if (lean_obj_tag(x_47) == 0) { lean_object* x_48; lean_object* x_49; @@ -3584,7 +3578,7 @@ lean_dec(x_72); x_75 = lean_ctor_get(x_73, 0); lean_inc(x_75); lean_dec(x_73); -x_76 = l_Lean_Kernel_isDiagnosticsEnabled(x_75); +x_76 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_75); lean_dec(x_75); if (x_76 == 0) { @@ -3783,7 +3777,7 @@ x_82 = lean_ctor_get(x_79, 0); x_83 = lean_ctor_get(x_79, 4); lean_dec(x_83); x_84 = l_Lean_getExternConstArityExport___closed__15; -x_85 = l_Lean_Kernel_enableDiag(x_82, x_84); +x_85 = l_Lean_Kernel_Environment_enableDiag(x_82, x_84); lean_ctor_set(x_79, 4, x_51); lean_ctor_set(x_79, 0, x_85); x_86 = lean_st_ref_set(x_70, x_79, x_80); @@ -3898,7 +3892,7 @@ lean_inc(x_109); lean_inc(x_108); lean_dec(x_79); x_115 = l_Lean_getExternConstArityExport___closed__15; -x_116 = l_Lean_Kernel_enableDiag(x_108, x_115); +x_116 = l_Lean_Kernel_Environment_enableDiag(x_108, x_115); x_117 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_117, 0, x_116); lean_ctor_set(x_117, 1, x_109); diff --git a/stage0/stdlib/Lean/Compiler/IR/ElimDeadBranches.c b/stage0/stdlib/Lean/Compiler/IR/ElimDeadBranches.c index fc5c9d485187..99c09d448a8d 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ElimDeadBranches.c +++ b/stage0/stdlib/Lean/Compiler/IR/ElimDeadBranches.c @@ -89,7 +89,7 @@ static lean_object* l___private_Lean_Compiler_IR_ElimDeadBranches_0__Lean_IR_Unr LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_UnreachableBranches_interpFnBody___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); lean_object* l_Std_Format_join(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_List_repr___at___private_Lean_Compiler_IR_ElimDeadBranches_0__Lean_IR_UnreachableBranches_reprValue____x40_Lean_Compiler_IR_ElimDeadBranches___hyg_43____spec__4___closed__7; static lean_object* l___private_Lean_Compiler_IR_ElimDeadBranches_0__Lean_IR_UnreachableBranches_reprValue____x40_Lean_Compiler_IR_ElimDeadBranches___hyg_43____closed__12; lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); @@ -3478,9 +3478,8 @@ x_14 = l_Lean_NameSet_contains(x_3, x_13); if (x_14 == 0) { lean_object* x_15; -lean_inc(x_13); lean_inc(x_1); -x_15 = lean_environment_find(x_1, x_13); +x_15 = l_Lean_Environment_find_x3f(x_1, x_13); if (lean_obj_tag(x_15) == 0) { size_t x_16; size_t x_17; lean_object* x_18; @@ -3648,9 +3647,8 @@ x_56 = l_Lean_NameSet_contains(x_3, x_55); if (x_56 == 0) { lean_object* x_57; -lean_inc(x_55); lean_inc(x_1); -x_57 = lean_environment_find(x_1, x_55); +x_57 = l_Lean_Environment_find_x3f(x_1, x_55); if (lean_obj_tag(x_57) == 0) { size_t x_58; size_t x_59; lean_object* x_60; lean_object* x_61; diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitC.c b/stage0/stdlib/Lean/Compiler/IR/EmitC.c index 3e0f4edd80b2..23bd97f2ba9e 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitC.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitC.c @@ -115,7 +115,7 @@ LEAN_EXPORT lean_object* l_List_forM___at_Lean_IR_EmitC_emitLns___spec__1___rarg uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitLns___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDeclAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_getJPParams(lean_object*, 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_IR_EmitC_emitFnDeclAux(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); @@ -3049,7 +3049,7 @@ x_7 = l_Lean_IR_EmitC_emitMainFn___lambda__1___closed__1; x_8 = lean_string_append(x_6, x_7); x_9 = l_Lean_IR_EmitC_emitLn___rarg___closed__1; x_10 = lean_string_append(x_8, x_9); -x_11 = lean_environment_find(x_1, x_2); +x_11 = l_Lean_Environment_find_x3f(x_1, x_2); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_3); @@ -3908,6 +3908,7 @@ lean_object* x_7; x_7 = l_Lean_IR_EmitC_emitMainFn___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_2); return x_7; } } @@ -3919,6 +3920,7 @@ x_7 = l_Lean_IR_EmitC_emitMainFn___lambda__2(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; } } @@ -3932,6 +3934,7 @@ x_9 = l_Lean_IR_EmitC_emitMainFn___lambda__3(x_1, x_2, x_3, x_8, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); +lean_dec(x_2); return x_9; } } @@ -3943,6 +3946,7 @@ x_6 = l_Lean_IR_EmitC_emitMainFn___lambda__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_6; } } diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c b/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c index 36bb111e3de7..109ebb439e89 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitLLVM.c @@ -178,7 +178,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitInitFn___lambda__1___boxed(lean_ LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_callLeanDecRef(size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); static lean_object* l_Lean_IR_EmitLLVM_emitDeclInit___closed__6; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitLLVM_emitArgSlot_____boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_IR_EmitLLVM_emitFnDeclAux___lambda__4(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -58477,7 +58477,6 @@ if (lean_obj_tag(x_80) == 0) uint8_t x_81; lean_dec(x_14); lean_dec(x_13); -lean_dec(x_11); lean_dec(x_10); x_81 = !lean_is_exclusive(x_79); if (x_81 == 0) @@ -58536,7 +58535,7 @@ lean_dec(x_80); x_91 = lean_ctor_get(x_79, 1); lean_inc(x_91); lean_dec(x_79); -x_92 = lean_environment_find(x_10, x_11); +x_92 = l_Lean_Environment_find_x3f(x_10, x_11); x_93 = l_Lean_IR_EmitLLVM_callLeanFinalizeTaskManager(x_1, x_2, x_13, x_14, x_91); if (lean_obj_tag(x_92) == 0) { @@ -58723,7 +58722,6 @@ else uint8_t x_145; lean_dec(x_14); lean_dec(x_13); -lean_dec(x_11); lean_dec(x_10); x_145 = !lean_is_exclusive(x_79); if (x_145 == 0) @@ -58890,7 +58888,6 @@ lean_dec(x_40); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); x_62 = !lean_is_exclusive(x_60); if (x_62 == 0) @@ -58980,7 +58977,6 @@ lean_dec(x_40); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); x_85 = !lean_is_exclusive(x_83); if (x_85 == 0) @@ -59178,7 +59174,6 @@ lean_dec(x_132); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); x_154 = lean_ctor_get(x_152, 1); lean_inc(x_154); @@ -59252,7 +59247,6 @@ lean_dec(x_132); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); x_173 = lean_ctor_get(x_171, 1); lean_inc(x_173); @@ -59455,7 +59449,6 @@ lean_dec(x_222); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); x_244 = lean_ctor_get(x_242, 1); lean_inc(x_244); @@ -59529,7 +59522,6 @@ lean_dec(x_222); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); x_263 = lean_ctor_get(x_261, 1); lean_inc(x_263); @@ -59670,7 +59662,6 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); x_14 = l___private_Init_Data_Repr_0__Nat_reprFast(x_9); x_15 = l_Lean_IR_EmitLLVM_emitMainFn___lambda__8___closed__1; x_16 = lean_string_append(x_15, x_14); @@ -59958,6 +59949,7 @@ x_22 = lean_unbox_usize(x_9); lean_dec(x_9); x_23 = l_Lean_IR_EmitLLVM_emitMainFn___lambda__6(x_16, x_17, x_18, x_19, x_5, x_20, x_21, x_8, x_22, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_12); +lean_dec(x_11); return x_23; } } @@ -59973,6 +59965,7 @@ x_12 = lean_unbox_usize(x_3); lean_dec(x_3); x_13 = l_Lean_IR_EmitLLVM_emitMainFn___lambda__7(x_10, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_6); +lean_dec(x_5); return x_13; } } @@ -59987,6 +59980,7 @@ lean_dec(x_2); x_11 = lean_unbox_usize(x_3); lean_dec(x_3); x_12 = l_Lean_IR_EmitLLVM_emitMainFn___lambda__8(x_9, x_10, x_11, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_4); return x_12; } } diff --git a/stage0/stdlib/Lean/Compiler/IR/UnboxResult.c b/stage0/stdlib/Lean/Compiler/IR/UnboxResult.c index 8785c2909aa5..ec87161ef74b 100644 --- a/stage0/stdlib/Lean/Compiler/IR/UnboxResult.c +++ b/stage0/stdlib/Lean/Compiler/IR/UnboxResult.c @@ -17,7 +17,7 @@ static lean_object* l_Lean_getConstInfo___at_Lean_IR_UnboxResult_initFn____x40_L static lean_object* l_Lean_getConstInfo___at_Lean_IR_UnboxResult_initFn____x40_Lean_Compiler_IR_UnboxResult___hyg_3____spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_IR_UnboxResult_initFn____x40_Lean_Compiler_IR_UnboxResult___hyg_3____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_UnboxResult_initFn____x40_Lean_Compiler_IR_UnboxResult___hyg_3_(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_IR_UnboxResult_initFn____x40_Lean_Compiler_IR_UnboxResult___hyg_3____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_UnboxResult_initFn____x40_Lean_Compiler_IR_UnboxResult___hyg_3____lambda__1___closed__4; @@ -137,8 +137,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -178,8 +177,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Compiler/ImplementedByAttr.c b/stage0/stdlib/Lean/Compiler/ImplementedByAttr.c index ddd4d15ba90a..715312b6dcab 100644 --- a/stage0/stdlib/Lean/Compiler/ImplementedByAttr.c +++ b/stage0/stdlib/Lean/Compiler/ImplementedByAttr.c @@ -25,7 +25,7 @@ static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAtt static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_3____closed__5; static lean_object* l_Lean_Compiler_getImplementedBy_x3f___closed__1; lean_object* l_Lean_ConstantInfo_name(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_3____lambda__4___closed__1; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_3____lambda__2___closed__4; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_3____closed__7; @@ -183,8 +183,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -224,8 +223,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Compiler/InitAttr.c b/stage0/stdlib/Lean/Compiler/InitAttr.c index d5cd21b8faca..8307db1b4ec6 100644 --- a/stage0/stdlib/Lean/Compiler/InitAttr.c +++ b/stage0/stdlib/Lean/Compiler/InitAttr.c @@ -46,7 +46,7 @@ LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_registerInitAttrUnsafe__ LEAN_EXPORT lean_object* l_Lean_isIOUnitInitFn___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getInitFnNameForCore_x3f___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1141____closed__23; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_registerInitAttrUnsafe___closed__1; lean_object* lean_run_mod_init(lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); @@ -156,6 +156,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1141____closed__9; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1141____closed__16; LEAN_EXPORT lean_object* l_Lean_registerInitAttrUnsafe___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_ParametricAttribute_getParam_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkAuxName___at_Lean_declareBuiltin___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -198,7 +199,6 @@ LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__2(lean_o LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_registerInitAttrUnsafe___spec__5(lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_declareBuiltin___closed__6; -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1141____closed__8; lean_object* l_Lean_Attribute_Builtin_getIdent_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -512,8 +512,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -553,8 +552,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -2583,7 +2581,7 @@ lean_inc(x_20); lean_dec(x_5); x_21 = lean_ctor_get(x_3, 0); lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 4); +x_22 = lean_ctor_get(x_21, 5); lean_inc(x_22); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); @@ -3826,7 +3824,7 @@ lean_dec(x_11); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); lean_dec(x_26); -x_28 = l_Lean_KernelException_toMessageData(x_27, x_13); +x_28 = l_Lean_Kernel_Exception_toMessageData(x_27, x_13); x_29 = l_Lean_MessageData_toString(x_28, x_23); x_30 = !lean_is_exclusive(x_29); if (x_30 == 0) @@ -4029,7 +4027,7 @@ lean_dec(x_11); x_92 = lean_ctor_get(x_91, 0); lean_inc(x_92); lean_dec(x_91); -x_93 = l_Lean_KernelException_toMessageData(x_92, x_13); +x_93 = l_Lean_Kernel_Exception_toMessageData(x_92, x_13); x_94 = l_Lean_MessageData_toString(x_93, x_88); x_95 = lean_ctor_get(x_94, 0); lean_inc(x_95); @@ -4193,7 +4191,7 @@ lean_dec(x_126); x_143 = lean_ctor_get(x_142, 0); lean_inc(x_143); lean_dec(x_142); -x_144 = l_Lean_KernelException_toMessageData(x_143, x_128); +x_144 = l_Lean_Kernel_Exception_toMessageData(x_143, x_128); x_145 = l_Lean_MessageData_toString(x_144, x_138); x_146 = lean_ctor_get(x_145, 0); lean_inc(x_146); diff --git a/stage0/stdlib/Lean/Compiler/InlineAttrs.c b/stage0/stdlib/Lean/Compiler/InlineAttrs.c index 8bc0c3f308aa..131c2731bf40 100644 --- a/stage0/stdlib/Lean/Compiler/InlineAttrs.c +++ b/stage0/stdlib/Lean/Compiler/InlineAttrs.c @@ -34,7 +34,7 @@ LEAN_EXPORT uint8_t lean_has_inline_if_reduce_attribute(lean_object*, lean_objec LEAN_EXPORT uint8_t l_Lean_Compiler_hasMacroInlineAttribute(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_InlineAttrs___hyg_267____closed__9; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_InlineAttrs___hyg_267____closed__3; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_inlineAttrs; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_InlineAttrs___hyg_267____closed__33; lean_object* l_Lean_EnumAttributes_setValue___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -448,8 +448,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -489,8 +488,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Check.c b/stage0/stdlib/Lean/Compiler/LCNF/Check.c index 5d18d61237fd..a61d5543cdfc 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Check.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Check.c @@ -55,7 +55,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_check uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqParam____x40_Lean_Compiler_LCNF_Basic___hyg_53_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Compiler_LCNF_Check_checkCases___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Check_checkFVar___closed__3; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Check_check___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Check_checkFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -8351,8 +8351,7 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_1); -x_15 = lean_environment_find(x_14, x_1); +x_15 = l_Lean_Environment_find_x3f(x_14, x_1); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -8392,8 +8391,7 @@ lean_dec(x_10); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_1); -x_27 = lean_environment_find(x_26, x_1); +x_27 = l_Lean_Environment_find_x3f(x_26, x_1); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/CompilerM.c b/stage0/stdlib/Lean/Compiler/LCNF/CompilerM.c index edf4148711ab..d7d60694aab1 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/CompilerM.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/CompilerM.c @@ -89,7 +89,7 @@ lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normExpr___rarg(uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normCodeImp___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instMonadFVarSubstNormalizerM(uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normCode(lean_object*); @@ -2875,7 +2875,8 @@ x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec(x_22); -x_24 = lean_environment_find(x_23, x_19); +x_24 = l_Lean_Environment_find_x3f(x_23, x_19); +lean_dec(x_19); if (lean_obj_tag(x_24) == 0) { uint8_t x_25; lean_object* x_26; @@ -2921,7 +2922,8 @@ lean_dec(x_20); x_34 = lean_ctor_get(x_32, 0); lean_inc(x_34); lean_dec(x_32); -x_35 = lean_environment_find(x_34, x_19); +x_35 = l_Lean_Environment_find_x3f(x_34, x_19); +lean_dec(x_19); if (lean_obj_tag(x_35) == 0) { uint8_t x_36; lean_object* x_37; lean_object* x_38; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c b/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c index a0fa14b99035..73a5ebe7e211 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ElimDeadBranches.c @@ -100,7 +100,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Unrea static lean_object* l_Lean_Compiler_LCNF_UnreachableBranches_addFunctionSummary___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_UnreachableBranches_findVarValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__7___closed__3; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_UnreachableBranches_Value_widening(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); @@ -2544,9 +2544,8 @@ x_13 = l_Lean_NameSet_contains(x_3, x_12); if (x_13 == 0) { lean_object* x_14; -lean_inc(x_12); lean_inc(x_1); -x_14 = lean_environment_find(x_1, x_12); +x_14 = l_Lean_Environment_find_x3f(x_1, x_12); if (lean_obj_tag(x_14) == 0) { size_t x_15; size_t x_16; lean_object* x_17; @@ -2711,9 +2710,8 @@ x_54 = l_Lean_NameSet_contains(x_3, x_53); if (x_54 == 0) { lean_object* x_55; -lean_inc(x_53); lean_inc(x_1); -x_55 = lean_environment_find(x_1, x_53); +x_55 = l_Lean_Environment_find_x3f(x_1, x_53); if (lean_obj_tag(x_55) == 0) { size_t x_56; size_t x_57; lean_object* x_58; lean_object* x_59; @@ -10731,8 +10729,7 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_35, 1); x_39 = lean_ctor_get(x_35, 0); lean_dec(x_39); -lean_inc(x_24); -x_40 = lean_environment_find(x_30, x_24); +x_40 = l_Lean_Environment_find_x3f(x_30, x_24); if (lean_obj_tag(x_40) == 0) { lean_object* x_41; @@ -10867,8 +10864,7 @@ lean_object* x_63; lean_object* x_64; x_63 = lean_ctor_get(x_35, 1); lean_inc(x_63); lean_dec(x_35); -lean_inc(x_24); -x_64 = lean_environment_find(x_30, x_24); +x_64 = l_Lean_Environment_find_x3f(x_30, x_24); if (lean_obj_tag(x_64) == 0) { lean_object* x_65; lean_object* x_66; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/InferType.c b/stage0/stdlib/Lean/Compiler/LCNF/InferType.c index 9002888e1071..2b6ca1b2c259 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/InferType.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/InferType.c @@ -72,7 +72,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_In LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_InferType_inferForallType_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_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_InferType_inferForallType_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_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_mkForallFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_InferType_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_mkFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_mkAuxFunDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -8405,8 +8405,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -8446,8 +8445,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; @@ -8977,7 +8975,8 @@ x_22 = lean_ctor_get(x_19, 1); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_find(x_23, x_17); +x_24 = l_Lean_Environment_find_x3f(x_23, x_17); +lean_dec(x_17); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; @@ -9807,7 +9806,8 @@ lean_dec(x_19); x_215 = lean_ctor_get(x_213, 0); lean_inc(x_215); lean_dec(x_213); -x_216 = lean_environment_find(x_215, x_17); +x_216 = l_Lean_Environment_find_x3f(x_215, x_17); +lean_dec(x_17); if (lean_obj_tag(x_216) == 0) { lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; @@ -10365,7 +10365,8 @@ if (lean_is_exclusive(x_341)) { x_345 = lean_ctor_get(x_342, 0); lean_inc(x_345); lean_dec(x_342); -x_346 = lean_environment_find(x_345, x_339); +x_346 = l_Lean_Environment_find_x3f(x_345, x_339); +lean_dec(x_339); if (lean_obj_tag(x_346) == 0) { lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Main.c b/stage0/stdlib/Lean/Compiler/LCNF/Main.c index ae905cd19f06..a5b8e5b4ba36 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Main.c @@ -93,10 +93,8 @@ lean_object* l_Lean_Compiler_LCNF_LCtx_toLocalContext(lean_object*); LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Compiler_LCNF_main___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_main___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_compile___closed__4; -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Compiler_LCNF_showDecl___closed__1; lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__5; @@ -129,6 +127,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_checkpo LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); lean_object* l_Lean_Compiler_LCNF_withPhase___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint64_t l_Lean_Compiler_LCNF_shouldGenerateCode___closed__2; @@ -261,6 +260,7 @@ static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__1; uint8_t l_Lean_isAuxRecursorWithSuffix(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__3; lean_object* lean_array_get_size(lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Compiler_LCNF_main___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_main___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -1700,7 +1700,7 @@ lean_dec(x_17); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = l_Lean_Kernel_isDiagnosticsEnabled(x_20); +x_21 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_20); lean_dec(x_20); if (x_21 == 0) { @@ -1750,7 +1750,7 @@ lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean x_27 = lean_ctor_get(x_24, 0); x_28 = lean_ctor_get(x_24, 4); lean_dec(x_28); -x_29 = l_Lean_Kernel_enableDiag(x_27, x_16); +x_29 = l_Lean_Kernel_Environment_enableDiag(x_27, x_16); x_30 = l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__8; lean_ctor_set(x_24, 4, x_30); lean_ctor_set(x_24, 0, x_29); @@ -1827,7 +1827,7 @@ lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); lean_dec(x_24); -x_52 = l_Lean_Kernel_enableDiag(x_45, x_16); +x_52 = l_Lean_Kernel_Environment_enableDiag(x_45, x_16); x_53 = l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__8; x_54 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_54, 0, x_52); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/PrettyPrinter.c b/stage0/stdlib/Lean/Compiler/LCNF/PrettyPrinter.c index 1bdca144cf14..c4c7f16bdb24 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/PrettyPrinter.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/PrettyPrinter.c @@ -67,10 +67,8 @@ static lean_object* l_Lean_Compiler_LCNF_PP_run___rarg___closed__3; static lean_object* l_Lean_Compiler_LCNF_PP_run___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PP_ppAlt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_PP_ppCode___closed__15; -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Compiler_LCNF_PP_ppExpr___closed__7; static lean_object* l_Lean_Compiler_LCNF_PP_ppAlt___closed__4; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Compiler_LCNF_PrettyPrinter_0__Lean_Compiler_LCNF_PP_prefixJoin___spec__1(lean_object*); static lean_object* l_Lean_Compiler_LCNF_runCompilerWithoutModifyingState___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PP_getFunType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -93,6 +91,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ppFunDecl(lean_object*, lean_objec static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Compiler_LCNF_PrettyPrinter_0__Lean_Compiler_LCNF_PP_join___spec__1___rarg___closed__3; LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_PP_ppFVar___lambda__1(lean_object*); extern lean_object* l_Lean_pp_all; +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_PrettyPrinter_0__Lean_Compiler_LCNF_PP_prefixJoin___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ppFunDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -183,6 +182,7 @@ lean_object* lean_string_append(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); extern lean_object* l_Lean_pp_letVarTypes; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Compiler_LCNF_PrettyPrinter_0__Lean_Compiler_LCNF_PP_prefixJoin___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Compiler_LCNF_runCompilerWithoutModifyingState___rarg___closed__2; static lean_object* l_Lean_Compiler_LCNF_PP_ppArg___closed__6; static uint64_t l_Lean_Compiler_LCNF_PP_ppExpr___closed__2; @@ -4988,7 +4988,7 @@ lean_dec(x_13); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -x_17 = l_Lean_Kernel_isDiagnosticsEnabled(x_16); +x_17 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_16); lean_dec(x_16); if (x_17 == 0) { @@ -5038,7 +5038,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean x_23 = lean_ctor_get(x_20, 0); x_24 = lean_ctor_get(x_20, 4); lean_dec(x_24); -x_25 = l_Lean_Kernel_enableDiag(x_23, x_12); +x_25 = l_Lean_Kernel_Environment_enableDiag(x_23, x_12); x_26 = l_Lean_Compiler_LCNF_PP_run___rarg___closed__3; lean_ctor_set(x_20, 4, x_26); lean_ctor_set(x_20, 0, x_25); @@ -5068,7 +5068,7 @@ lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); lean_dec(x_20); -x_38 = l_Lean_Kernel_enableDiag(x_31, x_12); +x_38 = l_Lean_Kernel_Environment_enableDiag(x_31, x_12); x_39 = l_Lean_Compiler_LCNF_PP_run___rarg___closed__3; x_40 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_40, 0, x_38); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/DiscrM.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/DiscrM.c index c220d2264e81..3346d29c08e8 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/DiscrM.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/DiscrM.c @@ -37,7 +37,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_findCtorName_x3f(lean_object* static lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_Simp_withDiscrCtorImp_updateCtx___spec__3___closed__1; static lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_Simp_withDiscrCtorImp_updateCtx___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Compiler_LCNF_Simp_simpCtorDiscrCore_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Compiler_LCNF_Simp_withDiscrCtorImp_updateCtx___spec__1___closed__1; static lean_object* l_Lean_getConstInfo___at_Lean_Compiler_LCNF_Simp_withDiscrCtorImp_updateCtx___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_Simp_withDiscrCtorImp_updateCtx___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -490,7 +490,8 @@ x_43 = lean_ctor_get(x_41, 0); x_44 = lean_ctor_get(x_43, 0); lean_inc(x_44); lean_dec(x_43); -x_45 = lean_environment_find(x_44, x_39); +x_45 = l_Lean_Environment_find_x3f(x_44, x_39); +lean_dec(x_39); if (lean_obj_tag(x_45) == 0) { lean_object* x_46; @@ -574,7 +575,8 @@ lean_dec(x_41); x_59 = lean_ctor_get(x_57, 0); lean_inc(x_59); lean_dec(x_57); -x_60 = lean_environment_find(x_59, x_39); +x_60 = l_Lean_Environment_find_x3f(x_59, x_39); +lean_dec(x_39); if (lean_obj_tag(x_60) == 0) { lean_object* x_61; lean_object* x_62; @@ -770,7 +772,8 @@ if (lean_is_exclusive(x_94)) { x_98 = lean_ctor_get(x_95, 0); lean_inc(x_98); lean_dec(x_95); -x_99 = lean_environment_find(x_98, x_92); +x_99 = l_Lean_Environment_find_x3f(x_98, x_92); +lean_dec(x_92); if (lean_obj_tag(x_99) == 0) { lean_object* x_100; lean_object* x_101; @@ -1520,8 +1523,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -1561,8 +1563,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/InlineProj.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/InlineProj.c index 3a47e9b011a3..42a32d1a9aef 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/InlineProj.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/InlineProj.c @@ -20,7 +20,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_mkLetDeclErased(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_Simp_inlineProjInst_x3f_visit___spec__1___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_inlineProjInst_x3f_visitCode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_inlineProjInst_x3f___lambda__1___closed__1; lean_object* l_Lean_Compiler_LCNF_Code_instantiateValueLevelParams_instCode(lean_object*, lean_object*, lean_object*); @@ -265,8 +265,7 @@ lean_dec(x_31); x_34 = lean_ctor_get(x_32, 0); lean_inc(x_34); lean_dec(x_32); -lean_inc(x_28); -x_35 = lean_environment_find(x_34, x_28); +x_35 = l_Lean_Environment_find_x3f(x_34, x_28); if (lean_obj_tag(x_35) == 0) { lean_object* x_36; lean_object* x_37; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c index 09f0dd7f2d35..7e673a1b48a1 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c @@ -56,7 +56,7 @@ lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParam___at_Lean_Compiler_LCNF_Simp_simpFunDecl___spec__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_etaPolyApp_x3f___closed__3; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParam___at_Lean_Compiler_LCNF_Simp_simpFunDecl___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___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_eraseFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -92,7 +92,6 @@ lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_norm lean_object* l_Lean_Compiler_LCNF_Simp_withInlining_check(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_eraseParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_isInductiveWithNoCtors(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f___closed__2; lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_addFVarSubst___spec__2(lean_object*); @@ -146,6 +145,7 @@ lean_object* l_Lean_Compiler_LCNF_Simp_inlineProjInst_x3f(lean_object*, lean_obj lean_object* l_Lean_Compiler_LCNF_Arg_toExpr(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at_Lean_Compiler_LCNF_Simp_simp___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* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instInhabitedParam; lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -1792,8 +1792,7 @@ if (lean_is_exclusive(x_21)) { x_25 = lean_ctor_get(x_22, 0); lean_inc(x_25); lean_dec(x_22); -lean_inc(x_17); -x_26 = lean_environment_find(x_25, x_17); +x_26 = l_Lean_Environment_find_x3f(x_25, x_17); if (lean_obj_tag(x_26) == 0) { lean_object* x_27; lean_object* x_28; @@ -9582,7 +9581,7 @@ lean_inc(x_1111); x_1118 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_1118, 0, x_1111); lean_ctor_set(x_1118, 1, x_1116); -x_1119 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_1117, x_1111, x_1113); +x_1119 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_1117, x_1111, x_1113); lean_ctor_set(x_3, 3, x_1119); lean_ctor_set(x_3, 2, x_1118); x_1120 = 0; @@ -11625,7 +11624,7 @@ lean_inc(x_1111); x_1483 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_1483, 0, x_1111); lean_ctor_set(x_1483, 1, x_1481); -x_1484 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_1482, x_1111, x_1113); +x_1484 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_1482, x_1111, x_1113); x_1485 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_1485, 0, x_1479); lean_ctor_set(x_1485, 1, x_1480); @@ -14942,7 +14941,7 @@ lean_inc(x_2050); x_2057 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_2057, 0, x_2050); lean_ctor_set(x_2057, 1, x_2055); -x_2058 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_2056, x_2050, x_2052); +x_2058 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_2056, x_2050, x_2052); lean_ctor_set(x_3, 3, x_2058); lean_ctor_set(x_3, 2, x_2057); x_2059 = l_Lean_Compiler_LCNF_Simp_specializePartialApp(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_2053); @@ -15185,7 +15184,7 @@ lean_inc(x_2050); x_2105 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_2105, 0, x_2050); lean_ctor_set(x_2105, 1, x_2103); -x_2106 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_2104, x_2050, x_2052); +x_2106 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_2104, x_2050, x_2052); x_2107 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_2107, 0, x_2101); lean_ctor_set(x_2107, 1, x_2102); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpM.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpM.c index 4aeb9c78f0a6..12f5da87089d 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpM.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpM.c @@ -69,7 +69,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(lean_ob LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_withAddMustInline(lean_object*); lean_object* l_Lean_Compiler_LCNF_LCtx_toLocalContext(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_addFunOcc(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Compiler_LCNF_Simp_FunDeclInfoMap_add___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_withInlining_check___closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_incInlineLocal___boxed(lean_object*); @@ -93,7 +92,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_withInlining_check(uint8_t, l lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_updateFunDeclInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_markSimplified___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_addFVarSubst___spec__2(lean_object*); uint8_t l_Lean_Compiler_LCNF_Code_sizeLe(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_incVisited___boxed(lean_object*); @@ -149,8 +147,10 @@ static lean_object* l_Lean_Compiler_LCNF_Simp_withIncRecDepth_throwMaxRecDepth__ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_withInlining_check___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_withInlining_check___lambda__2___closed__2; lean_object* l_Lean_Compiler_LCNF_Arg_toExpr(lean_object*); +lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_instMonadSimpM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); +lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_Simp_withIncRecDepth_throwMaxRecDepth___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_instMonadSimpM___closed__7; lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -2776,7 +2776,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_withInlining_check___lambda__ lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_ctor_get(x_4, 3); lean_inc(x_12); -x_13 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_12, x_1); +x_13 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_12, x_1); lean_inc(x_1); x_14 = l_Lean_Compiler_LCNF_getDecl_x3f(x_1, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_13) == 0) @@ -3093,7 +3093,7 @@ lean_inc(x_12); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_12); lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_18, x_12, x_14); +x_20 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_18, x_12, x_14); lean_ctor_set(x_4, 3, x_20); lean_ctor_set(x_4, 2, x_19); x_21 = lean_apply_8(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); @@ -3115,7 +3115,7 @@ lean_inc(x_12); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_12); lean_ctor_set(x_26, 1, x_24); -x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_25, x_12, x_14); +x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_25, x_12, x_14); x_28 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_28, 0, x_22); lean_ctor_set(x_28, 1, x_23); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpValue.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpValue.c index 0d36d26c820d..079261935de8 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpValue.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpValue.c @@ -18,7 +18,7 @@ static lean_object* l_Lean_Compiler_LCNF_Simp_applyImplementedBy_x3f___closed__1 LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simpAppApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Arg_toLetValue(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simpAppApp_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simpProj_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_applyImplementedBy_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -920,7 +920,8 @@ x_14 = lean_ctor_get(x_11, 1); x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); lean_dec(x_13); -x_16 = lean_environment_find(x_15, x_10); +x_16 = l_Lean_Environment_find_x3f(x_15, x_10); +lean_dec(x_10); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; @@ -1097,7 +1098,8 @@ lean_dec(x_11); x_52 = lean_ctor_get(x_50, 0); lean_inc(x_52); lean_dec(x_50); -x_53 = lean_environment_find(x_52, x_10); +x_53 = l_Lean_Environment_find_x3f(x_52, x_10); +lean_dec(x_10); if (lean_obj_tag(x_53) == 0) { lean_object* x_54; lean_object* x_55; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ToDecl.c b/stage0/stdlib/Lean/Compiler/LCNF/ToDecl.c index ad494f58bf92..d76ba83dd599 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ToDecl.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ToDecl.c @@ -45,7 +45,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Compiler_LC lean_object* l_Lean_Meta_Match_MatcherInfo_arity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_inlineMatchers___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_ToDecl_0__Lean_Compiler_LCNF_normalizeAlt___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_inlineMatchers___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_inlineAttrs; lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Meta_reduceMatcher_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2438,11 +2438,13 @@ x_9 = l_Lean_Compiler_LCNF_getDeclInfo_x3f___closed__1; lean_inc(x_1); x_10 = l_Lean_Name_str___override(x_1, x_9); lean_inc(x_8); -x_11 = lean_environment_find(x_8, x_10); +x_11 = l_Lean_Environment_find_x3f(x_8, x_10); +lean_dec(x_10); if (lean_obj_tag(x_11) == 0) { lean_object* x_12; -x_12 = lean_environment_find(x_8, x_1); +x_12 = l_Lean_Environment_find_x3f(x_8, x_1); +lean_dec(x_1); lean_ctor_set(x_5, 0, x_12); return x_5; } @@ -2485,11 +2487,13 @@ x_19 = l_Lean_Compiler_LCNF_getDeclInfo_x3f___closed__1; lean_inc(x_1); x_20 = l_Lean_Name_str___override(x_1, x_19); lean_inc(x_18); -x_21 = lean_environment_find(x_18, x_20); +x_21 = l_Lean_Environment_find_x3f(x_18, x_20); +lean_dec(x_20); if (lean_obj_tag(x_21) == 0) { lean_object* x_22; lean_object* x_23; -x_22 = lean_environment_find(x_18, x_1); +x_22 = l_Lean_Environment_find_x3f(x_18, x_1); +lean_dec(x_1); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_17); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ToLCNF.c b/stage0/stdlib/Lean/Compiler/LCNF/ToLCNF.c index f8aced264539..f508e6c054e9 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ToLCNF.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ToLCNF.c @@ -112,7 +112,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Compiler_LCN uint8_t lean_expr_has_loose_bvar(lean_object*, lean_object*); static lean_object* l_Lean_throwError___at_Lean_Compiler_LCNF_ToLCNF_bindCases_go___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToLCNF_toLCNF_visitAlt___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* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_findLetValue_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -7993,7 +7993,8 @@ lean_object* x_12; lean_object* x_13; x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -x_13 = lean_environment_find(x_1, x_12); +x_13 = l_Lean_Environment_find_x3f(x_1, x_12); +lean_dec(x_12); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -11583,9 +11584,8 @@ lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_3, 0); lean_inc(x_4); lean_dec(x_3); -lean_inc(x_4); lean_inc(x_1); -x_5 = lean_environment_find(x_1, x_4); +x_5 = l_Lean_Environment_find_x3f(x_1, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; uint8_t x_7; @@ -17080,8 +17080,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -17121,8 +17120,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c b/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c index 5e6bb169bddb..1ea851eab04a 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/ToMono.c @@ -40,7 +40,7 @@ static lean_object* l_Lean_Compiler_LCNF_trivialStructToMono___closed__12; lean_object* l_Lean_Compiler_LCNF_Arg_toLetValue(lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ToMono___hyg_2107____closed__5; lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_toMono(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_ToMono___hyg_2107____closed__14; @@ -290,7 +290,7 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -x_14 = lean_environment_find(x_13, x_1); +x_14 = l_Lean_Environment_find_x3f(x_13, x_1); if (lean_obj_tag(x_14) == 0) { lean_object* x_15; @@ -555,7 +555,7 @@ lean_dec(x_9); x_78 = lean_ctor_get(x_76, 0); lean_inc(x_78); lean_dec(x_76); -x_79 = lean_environment_find(x_78, x_1); +x_79 = l_Lean_Environment_find_x3f(x_78, x_1); if (lean_obj_tag(x_79) == 0) { lean_object* x_80; lean_object* x_81; @@ -743,6 +743,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_9; } } @@ -1694,7 +1695,6 @@ if (x_77 == 0) lean_object* x_78; lean_inc(x_6); lean_inc(x_5); -lean_inc(x_69); x_78 = l_Lean_Compiler_LCNF_isTrivialConstructorApp_x3f(x_69, x_70, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_78) == 0) { @@ -1716,8 +1716,7 @@ lean_dec(x_81); x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); lean_dec(x_82); -lean_inc(x_69); -x_85 = lean_environment_find(x_84, x_69); +x_85 = l_Lean_Environment_find_x3f(x_84, x_69); if (lean_obj_tag(x_85) == 0) { size_t x_86; size_t x_87; lean_object* x_88; uint8_t x_89; @@ -1948,7 +1947,6 @@ if (x_137 == 0) lean_object* x_138; lean_inc(x_6); lean_inc(x_5); -lean_inc(x_130); x_138 = l_Lean_Compiler_LCNF_isTrivialConstructorApp_x3f(x_130, x_131, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_138) == 0) { @@ -1970,8 +1968,7 @@ lean_dec(x_141); x_144 = lean_ctor_get(x_142, 0); lean_inc(x_144); lean_dec(x_142); -lean_inc(x_130); -x_145 = lean_environment_find(x_144, x_130); +x_145 = l_Lean_Environment_find_x3f(x_144, x_130); if (lean_obj_tag(x_145) == 0) { size_t x_146; size_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Types.c b/stage0/stdlib/Lean/Compiler/LCNF/Types.c index 953d81b72b75..ef89b58607e1 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Types.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Types.c @@ -38,7 +38,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_isPropFormerType_go___lambda__1(le lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_erasedExpr; lean_object* lean_array_fget(lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instantiateForall_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); @@ -3034,7 +3034,8 @@ x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); lean_dec(x_9); -x_11 = lean_environment_find(x_10, x_6); +x_11 = l_Lean_Environment_find_x3f(x_10, x_6); +lean_dec(x_6); if (lean_obj_tag(x_11) == 0) { uint8_t x_12; lean_object* x_13; @@ -3086,7 +3087,8 @@ lean_dec(x_7); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_find(x_24, x_6); +x_25 = l_Lean_Environment_find_x3f(x_24, x_6); +lean_dec(x_6); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_object* x_27; lean_object* x_28; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Util.c b/stage0/stdlib/Lean/Compiler/LCNF/Util.c index ac553e1c1774..b35eeb965886 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Util.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Util.c @@ -33,7 +33,7 @@ extern lean_object* l_Lean_casesOnSuffix; static lean_object* l_Lean_Compiler_LCNF_isLcCast_x3f___closed__2; static lean_object* l_Lean_Compiler_LCNF_isCasesApp_x3f___closed__4; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__26; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_getCasesInfo_x3f___spec__2___closed__3; static lean_object* l_Lean_Compiler_LCNF_builtinRuntimeTypes___closed__36; lean_object* l_Lean_stringToMessageData(lean_object*); @@ -297,8 +297,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -338,8 +337,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Compiler/Old.c b/stage0/stdlib/Lean/Compiler/Old.c index 8159c99833fe..14bbc560e0a2 100644 --- a/stage0/stdlib/Lean/Compiler/Old.c +++ b/stage0/stdlib/Lean/Compiler/Old.c @@ -18,9 +18,8 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_checkIsDefinition___lambda__1___boxed(l lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Compiler_mkUnsafeRecName___closed__1; static lean_object* l_Lean_Compiler_mkEagerLambdaLiftingName___closed__1; -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Compiler_getDeclNamesForCodeGen___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_eager_lambda_lifting_name(lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_getDeclNamesForCodeGen(lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_checkIsDefinition___lambda__1(lean_object*); @@ -30,6 +29,7 @@ static lean_object* l_Lean_Compiler_isEagerLambdaLiftingName___closed__1; lean_object* lean_compile_decls(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_compileDecls___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Declaration_getNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_checkIsDefinition___closed__5; static lean_object* l_Lean_Compiler_checkIsDefinition___closed__4; LEAN_EXPORT lean_object* lean_is_unsafe_rec_name(lean_object*); @@ -38,7 +38,6 @@ static lean_object* l_Lean_Compiler_checkIsDefinition___closed__2; uint8_t l_String_isPrefixOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_checkIsDefinition(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_unsafe_rec_name(lean_object*); -lean_object* l_List_reverse___rarg(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_is_eager_lambda_lifting_name(lean_object*); static lean_object* l_Lean_Compiler_checkIsDefinition___closed__3; @@ -126,62 +125,6 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Compiler_getDeclNamesForCodeGen___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -lean_dec(x_4); -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_1, 1); -x_8 = lean_ctor_get(x_1, 0); -lean_dec(x_8); -x_9 = lean_ctor_get(x_5, 0); -lean_inc(x_9); -lean_dec(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_9); -{ -lean_object* _tmp_0 = x_7; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); -lean_dec(x_1); -x_12 = lean_ctor_get(x_5, 0); -lean_inc(x_12); -lean_dec(x_5); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_2); -x_1 = x_11; -x_2 = x_13; -goto _start; -} -} -} -} LEAN_EXPORT lean_object* l_Lean_Compiler_getDeclNamesForCodeGen(lean_object* x_1) { _start: { @@ -206,7 +149,7 @@ x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); lean_dec(x_1); x_5 = lean_box(0); -x_6 = l_List_mapTR_loop___at_Lean_Compiler_getDeclNamesForCodeGen___spec__1(x_4, x_5); +x_6 = l_List_mapTR_loop___at_Lean_Declaration_getNames___spec__1(x_4, x_5); return x_6; } case 6: @@ -291,8 +234,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_checkIsDefinition(lean_object* x_1, lea _start: { lean_object* x_3; -lean_inc(x_2); -x_3 = lean_environment_find(x_1, x_2); +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; diff --git a/stage0/stdlib/Lean/CoreM.c b/stage0/stdlib/Lean/CoreM.c index 5da8820315a8..ea5733dc14f6 100644 --- a/stage0/stdlib/Lean/CoreM.c +++ b/stage0/stdlib/Lean/CoreM.c @@ -256,11 +256,9 @@ static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_185____closed__15 LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadRefCoreM___closed__1; lean_object* l_Lean_Environment_compileDecl(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__15; LEAN_EXPORT lean_object* l_Lean_Core_CoreM_run_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___at_Lean_Core_wrapAsyncAsSnapshot___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadOptionsCoreM(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_3066____closed__5; @@ -358,6 +356,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Core_insta LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l___auto____x40_Lean_CoreM___hyg_4099____closed__5; LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported(lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__4; @@ -577,12 +576,13 @@ static lean_object* l_Lean_compileDecl___closed__1; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_185____closed__7; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_ImportM_runCoreM___rarg___closed__8; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_addTraceAsMessages___spec__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_useDiagnosticMsg___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadRuntimeExceptionStateRefT_x27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_wrapAsyncAsSnapshot___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instFunctorOfMonad___rarg(lean_object*); static lean_object* l___auto____x40_Lean_CoreM___hyg_4099____closed__37; @@ -753,6 +753,7 @@ static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_4055____closed__1 static lean_object* l_Lean_addTraceAsMessages___at_Lean_Core_wrapAsyncAsSnapshot___spec__6___lambda__1___closed__5; static lean_object* l___auto____x40_Lean_CoreM___hyg_4099____closed__33; LEAN_EXPORT lean_object* l_Lean_Core_instAddMessageContextCoreM; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_instMonadExceptOfExceptionCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_185____closed__9; static lean_object* l_Lean_useDiagnosticMsg___lambda__2___closed__5; @@ -800,7 +801,6 @@ LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM___lambda__2(lean_object*, LEAN_EXPORT lean_object* l_Lean_Core_wrapAsyncAsSnapshot___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_lazy(lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Core_wrapAsyncAsSnapshot___spec__1___lambda__3___closed__2; -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadRuntimeExceptionReaderT(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_maxHeartbeats; static lean_object* l_Lean_Core_instMonadResolveNameCoreM___closed__3; @@ -2165,7 +2165,7 @@ lean_dec(x_10); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -x_14 = l_Lean_Kernel_isDiagnosticsEnabled(x_13); +x_14 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_13); lean_dec(x_13); if (x_14 == 0) { @@ -2219,7 +2219,7 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean x_20 = lean_ctor_get(x_17, 0); x_21 = lean_ctor_get(x_17, 4); lean_dec(x_21); -x_22 = l_Lean_Kernel_enableDiag(x_20, x_9); +x_22 = l_Lean_Kernel_Environment_enableDiag(x_20, x_9); x_23 = l_Lean_Core_instInhabitedCache___closed__3; lean_ctor_set(x_17, 4, x_23); lean_ctor_set(x_17, 0, x_22); @@ -2249,7 +2249,7 @@ lean_inc(x_30); lean_inc(x_29); lean_inc(x_28); lean_dec(x_17); -x_35 = l_Lean_Kernel_enableDiag(x_28, x_9); +x_35 = l_Lean_Kernel_Environment_enableDiag(x_28, x_9); x_36 = l_Lean_Core_instInhabitedCache___closed__3; x_37 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_37, 0, x_35); @@ -2315,7 +2315,7 @@ lean_dec(x_8); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -x_12 = l_Lean_Kernel_isDiagnosticsEnabled(x_11); +x_12 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_11); lean_dec(x_11); if (x_12 == 0) { @@ -2369,7 +2369,7 @@ lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean x_18 = lean_ctor_get(x_15, 0); x_19 = lean_ctor_get(x_15, 4); lean_dec(x_19); -x_20 = l_Lean_Kernel_enableDiag(x_18, x_7); +x_20 = l_Lean_Kernel_Environment_enableDiag(x_18, x_7); x_21 = l_Lean_Core_instInhabitedCache___closed__3; lean_ctor_set(x_15, 4, x_21); lean_ctor_set(x_15, 0, x_20); @@ -2399,7 +2399,7 @@ lean_inc(x_28); lean_inc(x_27); lean_inc(x_26); lean_dec(x_15); -x_33 = l_Lean_Kernel_enableDiag(x_26, x_7); +x_33 = l_Lean_Kernel_Environment_enableDiag(x_26, x_7); x_34 = l_Lean_Core_instInhabitedCache___closed__3; x_35 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_35, 0, x_33); @@ -3214,7 +3214,8 @@ x_6 = lean_ctor_get(x_4, 0); x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); lean_dec(x_6); -x_8 = lean_environment_main_module(x_7); +x_8 = l_Lean_Environment_mainModule(x_7); +lean_dec(x_7); lean_ctor_set(x_4, 0, x_8); return x_4; } @@ -3229,7 +3230,8 @@ lean_dec(x_4); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -x_12 = lean_environment_main_module(x_11); +x_12 = l_Lean_Environment_mainModule(x_11); +lean_dec(x_11); x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_10); @@ -6618,7 +6620,8 @@ x_16 = lean_ctor_get(x_14, 0); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); lean_dec(x_16); -x_18 = lean_environment_main_module(x_17); +x_18 = l_Lean_Environment_mainModule(x_17); +lean_dec(x_17); x_19 = l_Lean_addMacroScope(x_18, x_1, x_9); lean_ctor_set(x_14, 0, x_19); return x_14; @@ -6634,7 +6637,8 @@ lean_dec(x_14); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_Lean_addMacroScope(x_23, x_1, x_9); x_25 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_25, 0, x_24); @@ -6693,7 +6697,8 @@ if (lean_is_exclusive(x_39)) { x_43 = lean_ctor_get(x_40, 0); lean_inc(x_43); lean_dec(x_40); -x_44 = lean_environment_main_module(x_43); +x_44 = l_Lean_Environment_mainModule(x_43); +lean_dec(x_43); x_45 = l_Lean_addMacroScope(x_44, x_1, x_27); if (lean_is_scalar(x_42)) { x_46 = lean_alloc_ctor(0, 2, 0); @@ -6764,7 +6769,7 @@ if (lean_is_exclusive(x_11)) { x_15 = lean_ctor_get(x_12, 0); lean_inc(x_15); lean_dec(x_12); -x_16 = l_Lean_Kernel_isDiagnosticsEnabled(x_15); +x_16 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_15); lean_dec(x_15); if (x_16 == 0) { @@ -6819,7 +6824,7 @@ lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean x_22 = lean_ctor_get(x_19, 0); x_23 = lean_ctor_get(x_19, 4); lean_dec(x_23); -x_24 = l_Lean_Kernel_enableDiag(x_22, x_10); +x_24 = l_Lean_Kernel_Environment_enableDiag(x_22, x_10); x_25 = l_Lean_Core_instInhabitedCache___closed__3; lean_ctor_set(x_19, 4, x_25); lean_ctor_set(x_19, 0, x_24); @@ -6983,7 +6988,7 @@ lean_inc(x_61); lean_inc(x_60); lean_inc(x_59); lean_dec(x_19); -x_66 = l_Lean_Kernel_enableDiag(x_59, x_10); +x_66 = l_Lean_Kernel_Environment_enableDiag(x_59, x_10); x_67 = l_Lean_Core_instInhabitedCache___closed__3; x_68 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_68, 0, x_66); @@ -7185,7 +7190,7 @@ lean_dec(x_11); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -x_15 = l_Lean_Kernel_isDiagnosticsEnabled(x_14); +x_15 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_14); lean_dec(x_14); if (x_15 == 0) { @@ -7239,7 +7244,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean x_21 = lean_ctor_get(x_18, 0); x_22 = lean_ctor_get(x_18, 4); lean_dec(x_22); -x_23 = l_Lean_Kernel_enableDiag(x_21, x_10); +x_23 = l_Lean_Kernel_Environment_enableDiag(x_21, x_10); x_24 = l_Lean_Core_instInhabitedCache___closed__3; lean_ctor_set(x_18, 4, x_24); lean_ctor_set(x_18, 0, x_23); @@ -7323,7 +7328,7 @@ lean_inc(x_42); lean_inc(x_41); lean_inc(x_40); lean_dec(x_18); -x_47 = l_Lean_Kernel_enableDiag(x_40, x_10); +x_47 = l_Lean_Kernel_Environment_enableDiag(x_40, x_10); x_48 = l_Lean_Core_instInhabitedCache___closed__3; x_49 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_49, 0, x_47); @@ -7513,7 +7518,7 @@ lean_dec(x_41); x_44 = lean_ctor_get(x_42, 0); lean_inc(x_44); lean_dec(x_42); -x_45 = l_Lean_Kernel_isDiagnosticsEnabled(x_44); +x_45 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_44); lean_dec(x_44); if (x_45 == 0) { @@ -7567,7 +7572,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_48, 0); x_52 = lean_ctor_get(x_48, 4); lean_dec(x_52); -x_53 = l_Lean_Kernel_enableDiag(x_51, x_40); +x_53 = l_Lean_Kernel_Environment_enableDiag(x_51, x_40); x_54 = l_Lean_Core_instInhabitedCache___closed__3; lean_ctor_set(x_48, 4, x_54); lean_ctor_set(x_48, 0, x_53); @@ -7679,7 +7684,7 @@ lean_inc(x_79); lean_inc(x_78); lean_inc(x_77); lean_dec(x_48); -x_84 = l_Lean_Kernel_enableDiag(x_77, x_40); +x_84 = l_Lean_Kernel_Environment_enableDiag(x_77, x_40); x_85 = l_Lean_Core_instInhabitedCache___closed__3; x_86 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_86, 0, x_84); @@ -7909,7 +7914,7 @@ lean_dec(x_146); x_149 = lean_ctor_get(x_147, 0); lean_inc(x_149); lean_dec(x_147); -x_150 = l_Lean_Kernel_isDiagnosticsEnabled(x_149); +x_150 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_149); lean_dec(x_149); if (x_150 == 0) { @@ -7984,7 +7989,7 @@ if (lean_is_exclusive(x_153)) { lean_dec_ref(x_153); x_162 = lean_box(0); } -x_163 = l_Lean_Kernel_enableDiag(x_155, x_145); +x_163 = l_Lean_Kernel_Environment_enableDiag(x_155, x_145); x_164 = l_Lean_Core_instInhabitedCache___closed__3; if (lean_is_scalar(x_162)) { x_165 = lean_alloc_ctor(0, 8, 0); @@ -10225,7 +10230,7 @@ lean_dec(x_11); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -x_15 = l_Lean_Kernel_isDiagnosticsEnabled(x_14); +x_15 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_14); lean_dec(x_14); if (x_15 == 0) { @@ -10279,7 +10284,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean x_21 = lean_ctor_get(x_18, 0); x_22 = lean_ctor_get(x_18, 4); lean_dec(x_22); -x_23 = l_Lean_Kernel_enableDiag(x_21, x_10); +x_23 = l_Lean_Kernel_Environment_enableDiag(x_21, x_10); x_24 = l_Lean_Core_instInhabitedCache___closed__3; lean_ctor_set(x_18, 4, x_24); lean_ctor_set(x_18, 0, x_23); @@ -10363,7 +10368,7 @@ lean_inc(x_42); lean_inc(x_41); lean_inc(x_40); lean_dec(x_18); -x_47 = l_Lean_Kernel_enableDiag(x_40, x_10); +x_47 = l_Lean_Kernel_Environment_enableDiag(x_40, x_10); x_48 = l_Lean_Core_instInhabitedCache___closed__3; x_49 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_49, 0, x_47); @@ -18761,7 +18766,6 @@ x_4 = lean_is_aux_recursor(x_1, x_3); if (x_4 == 0) { uint8_t x_5; -lean_inc(x_3); x_5 = l_Lean_isRecCore(x_1, x_3); if (x_5 == 0) { @@ -18820,7 +18824,6 @@ return x_16; else { uint8_t x_17; -lean_inc(x_3); x_17 = l_Lean_isRecCore(x_1, x_3); if (x_17 == 0) { @@ -21510,7 +21513,7 @@ LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_compileDecl___spe lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_2, 2); lean_inc(x_5); -x_6 = l_Lean_KernelException_toMessageData(x_1, x_5); +x_6 = l_Lean_Kernel_Exception_toMessageData(x_1, x_5); x_7 = l_Lean_throwError___at_Lean_compileDecl___spec__5(x_6, x_2, x_3, x_4); lean_dec(x_2); return x_7; @@ -23501,7 +23504,7 @@ lean_dec(x_17); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = l_Lean_Kernel_isDiagnosticsEnabled(x_20); +x_21 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_20); lean_dec(x_20); if (x_21 == 0) { @@ -23555,7 +23558,7 @@ lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean x_27 = lean_ctor_get(x_24, 0); x_28 = lean_ctor_get(x_24, 4); lean_dec(x_28); -x_29 = l_Lean_Kernel_enableDiag(x_27, x_16); +x_29 = l_Lean_Kernel_Environment_enableDiag(x_27, x_16); lean_ctor_set(x_24, 4, x_5); lean_ctor_set(x_24, 0, x_29); x_30 = lean_st_ref_set(x_8, x_24, x_25); @@ -23584,7 +23587,7 @@ lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_dec(x_24); -x_41 = l_Lean_Kernel_enableDiag(x_34, x_16); +x_41 = l_Lean_Kernel_Environment_enableDiag(x_34, x_16); x_42 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_35); @@ -23666,7 +23669,7 @@ lean_dec(x_70); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = l_Lean_Kernel_isDiagnosticsEnabled(x_73); +x_74 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_73); lean_dec(x_73); if (x_74 == 0) { @@ -23741,7 +23744,7 @@ if (lean_is_exclusive(x_77)) { lean_dec_ref(x_77); x_86 = lean_box(0); } -x_87 = l_Lean_Kernel_enableDiag(x_79, x_69); +x_87 = l_Lean_Kernel_Environment_enableDiag(x_79, x_69); if (lean_is_scalar(x_86)) { x_88 = lean_alloc_ctor(0, 8, 0); } else { @@ -23952,7 +23955,7 @@ lean_dec(x_70); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = l_Lean_Kernel_isDiagnosticsEnabled(x_73); +x_74 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_73); lean_dec(x_73); if (x_74 == 0) { @@ -24142,7 +24145,7 @@ x_80 = lean_ctor_get(x_77, 0); x_81 = lean_ctor_get(x_77, 4); lean_dec(x_81); x_82 = l_Lean_ImportM_runCoreM___rarg___closed__10; -x_83 = l_Lean_Kernel_enableDiag(x_80, x_82); +x_83 = l_Lean_Kernel_Environment_enableDiag(x_80, x_82); lean_ctor_set(x_77, 4, x_49); lean_ctor_set(x_77, 0, x_83); x_84 = lean_st_ref_set(x_68, x_77, x_78); @@ -24257,7 +24260,7 @@ lean_inc(x_107); lean_inc(x_106); lean_dec(x_77); x_113 = l_Lean_ImportM_runCoreM___rarg___closed__10; -x_114 = l_Lean_Kernel_enableDiag(x_106, x_113); +x_114 = l_Lean_Kernel_Environment_enableDiag(x_106, x_113); x_115 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_115, 0, x_114); lean_ctor_set(x_115, 1, x_107); diff --git a/stage0/stdlib/Lean/Data/Json/Printer.c b/stage0/stdlib/Lean/Data/Json/Printer.c index 39b0a4548a8a..8131e3c14116 100644 --- a/stage0/stdlib/Lean/Data/Json/Printer.c +++ b/stage0/stdlib/Lean/Data/Json/Printer.c @@ -13,89 +13,3713 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111; lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141; LEAN_EXPORT lean_object* l_Lean_Json_compress(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94; uint32_t lean_string_utf8_get(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181; +lean_object* lean_byte_array_mk(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114; static lean_object* l_Lean_Json_compress_go___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137; lean_object* l_Lean_JsonNumber_toString(lean_object*); +uint8_t lean_byte_array_fget(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183; lean_object* lean_uint32_to_nat(uint32_t); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_Json_compress_go___spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47; LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__3; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62; static lean_object* l_Lean_Json_render___closed__9; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168; static lean_object* l_Lean_Json_render___closed__19; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21; uint8_t lean_usize_dec_eq(size_t, size_t); +uint8_t lean_string_get_byte_fast(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27; static lean_object* l_Lean_Json_renderString___closed__1; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__2; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103; static lean_object* l_Lean_Json_render___closed__12; lean_object* lean_string_utf8_byte_size(lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); static lean_object* l_Lean_Json_render___closed__20; static lean_object* l_Lean_RBNode_fold___at_Lean_Json_render___spec__2___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101; static lean_object* l_Lean_Json_render___closed__6; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78; +LEAN_EXPORT uint8_t l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105; size_t lean_usize_of_nat(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167; static lean_object* l_Lean_Json_render___closed__14; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240; lean_object* lean_string_utf8_next(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204; LEAN_EXPORT lean_object* l_Lean_Json_escape(lean_object*, lean_object*); uint32_t l_Nat_digitChar(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Json_compress_go___spec__2(lean_object*, size_t, size_t, lean_object*); uint8_t lean_uint32_dec_le(uint32_t, uint32_t); lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163; lean_object* lean_nat_div(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234; static lean_object* l_Lean_Json_render___closed__17; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238; static lean_object* l_Lean_Json_render___closed__4; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Json_compress_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42; static lean_object* l_Lean_Json_render___closed__11; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79; lean_object* lean_array_to_list(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Json_compress_go___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Json_render___closed__16; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133; static lean_object* l_Lean_Json_render___closed__13; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131; LEAN_EXPORT lean_object* l_Lean_Json_renderString___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71; static lean_object* l_Lean_Json_instToFormat___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__5; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Json_render___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81; static lean_object* l_Lean_Json_render___closed__5; static lean_object* l_Lean_Json_render___closed__7; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148; static lean_object* l_Lean_Json_render___closed__15; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92; static lean_object* l_Lean_Json_render___closed__3; static lean_object* l_Lean_Json_render___closed__10; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70; LEAN_EXPORT lean_object* l_String_foldlAux___at_Lean_Json_escape___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237; static lean_object* l_Lean_Json_render___closed__1; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211; +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76; lean_object* lean_string_length(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Json_compress_go___spec__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122; lean_object* lean_nat_mod(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4; LEAN_EXPORT lean_object* l_Lean_Json_instToString(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Json_render___spec__1(size_t, size_t, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_Json_compress_go___spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54; lean_object* l_Std_Format_joinSep___at_Prod_repr___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91; LEAN_EXPORT lean_object* l_String_foldlAux___at_Lean_Json_escape___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_renderString(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_compress_go(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232; static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__4; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129; size_t lean_usize_sub(size_t, size_t); +lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164; +lean_object* lean_uint8_to_nat(uint8_t); size_t lean_usize_add(size_t, size_t); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104; static lean_object* l_Lean_Json_render___closed__18; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32; LEAN_EXPORT lean_object* l_Lean_Json_escape___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165; static lean_object* l_Lean_Json_render___closed__8; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5; lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_render(lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60; static lean_object* l_Lean_Json_render___closed__2; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_RBNode_fold___at_Lean_Json_render___spec__2___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143; +lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Json_render___closed__21; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_Json_render___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_instToFormat; LEAN_EXPORT lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16; +uint8_t lean_uint8_dec_eq(uint8_t, uint8_t); +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux(lean_object*, uint32_t); static lean_object* l_Lean_RBNode_fold___at_Lean_Json_render___spec__2___closed__3; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255; +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210; +static lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127; +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = lean_box(0); +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128; +x_2 = 0; +x_3 = lean_box(x_2); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_1); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257; +x_2 = lean_byte_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258; +return x_1; +} +} static lean_object* _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1() { _start: { @@ -265,6 +3889,79 @@ x_4 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux(x_1, x_3); return x_4; } } +LEAN_EXPORT uint8_t l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_string_utf8_byte_size(x_1); +x_4 = lean_nat_dec_lt(x_2, x_3); +lean_dec(x_3); +if (x_4 == 0) +{ +uint8_t x_5; +lean_dec(x_2); +x_5 = 0; +return x_5; +} +else +{ +uint8_t x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; uint8_t x_11; +lean_inc(x_2); +x_6 = lean_string_get_byte_fast(x_1, x_2); +x_7 = lean_uint8_to_nat(x_6); +x_8 = l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable; +x_9 = lean_byte_array_fget(x_8, x_7); +lean_dec(x_7); +x_10 = 0; +x_11 = lean_uint8_dec_eq(x_9, x_10); +if (x_11 == 0) +{ +uint8_t x_12; +lean_dec(x_2); +x_12 = 1; +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_2, x_13); +lean_dec(x_2); +x_2 = x_14; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(x_1, x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_String_foldlAux___at_Lean_Json_escape___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -291,13 +3988,24 @@ goto _start; LEAN_EXPORT lean_object* l_Lean_Json_escape(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = lean_string_utf8_byte_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_String_foldlAux___at_Lean_Json_escape___spec__1(x_1, x_3, x_4, x_2); -lean_dec(x_3); +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l___private_Lean_Data_Json_Printer_0__Lean_Json_needEscape_go(x_1, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = lean_string_append(x_2, x_1); return x_5; } +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_string_utf8_byte_size(x_1); +x_7 = l_String_foldlAux___at_Lean_Json_escape___spec__1(x_1, x_6, x_3, x_2); +lean_dec(x_6); +return x_7; +} +} } LEAN_EXPORT lean_object* l_String_foldlAux___at_Lean_Json_escape___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: @@ -2030,6 +5738,524 @@ lean_dec_ref(res); res = initialize_Init_Data_List_Impl(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__1); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__2); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__3); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__4); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__5); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__6); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__7); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__8); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__9); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__10); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__11); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__12); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__13); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__14); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__15); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__16); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__17); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__18); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__19); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__20); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__21); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__22); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__23); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__24); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__25); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__26); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__27); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__28); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__29); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__30); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__31); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__32); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__33); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__34); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__35); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__36); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__37); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__38); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__39); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__40); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__41); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__42); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__43); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__44); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__45); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__46); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__47); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__48); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__49); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__50); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__51); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__52); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__53); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__54); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__55); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__56); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__57); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__58); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__59); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__60); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__61); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__62); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__63); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__64); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__65); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__66); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__67); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__68); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__69); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__70); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__71); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__72); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__73); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__74); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__75); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__76); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__77); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__78); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__79); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__80); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__81); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__82); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__83); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__84); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__85); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__86); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__87); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__88); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__89); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__90); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__91); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__92); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__93); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__94); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__95); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__96); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__97); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__98); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__99); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__100); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__101); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__102); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__103); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__104); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__105); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__106); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__107); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__108); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__109); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__110); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__111); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__112); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__113); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__114); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__115); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__116); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__117); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__118); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__119); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__120); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__121); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__122); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__123); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__124); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__125); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__126); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__127); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__128); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__129); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__130); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__131); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__132); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__133); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__134); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__135); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__136); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__137); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__138); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__139); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__140); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__141); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__142); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__143); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__144); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__145); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__146); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__147); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__148); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__149); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__150); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__151); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__152); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__153); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__154); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__155); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__156); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__157); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__158); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__159); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__160); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__161); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__162); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__163); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__164); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__165); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__166); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__167); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__168); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__169); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__170); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__171); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__172); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__173); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__174); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__175); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__176); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__177); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__178); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__179); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__180); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__181); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__182); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__183); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__184); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__185); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__186); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__187); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__188); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__189); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__190); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__191); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__192); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__193); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__194); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__195); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__196); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__197); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__198); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__199); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__200); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__201); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__202); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__203); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__204); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__205); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__206); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__207); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__208); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__209); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__210); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__211); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__212); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__213); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__214); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__215); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__216); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__217); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__218); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__219); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__220); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__221); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__222); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__223); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__224); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__225); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__226); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__227); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__228); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__229); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__230); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__231); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__232); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__233); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__234); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__235); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__236); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__237); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__238); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__239); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__240); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__241); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__242); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__243); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__244); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__245); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__246); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__247); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__248); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__249); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__250); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__251); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__252); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__253); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__254); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__255); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__256); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__257); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable___closed__258); +l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable(); +lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeTable); l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1(); lean_mark_persistent(l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__1); l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__2 = _init_l___private_Lean_Data_Json_Printer_0__Lean_Json_escapeAux___closed__2(); diff --git a/stage0/stdlib/Lean/Data/Lsp.c b/stage0/stdlib/Lean/Data/Lsp.c index 0055ee3f3b15..bd31f38bcb5e 100644 --- a/stage0/stdlib/Lean/Data/Lsp.c +++ b/stage0/stdlib/Lean/Data/Lsp.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Data.Lsp -// Imports: Lean.Data.Lsp.Basic Lean.Data.Lsp.Capabilities Lean.Data.Lsp.Client Lean.Data.Lsp.Communication Lean.Data.Lsp.Diagnostics Lean.Data.Lsp.Extra Lean.Data.Lsp.InitShutdown Lean.Data.Lsp.Internal Lean.Data.Lsp.LanguageFeatures Lean.Data.Lsp.TextSync Lean.Data.Lsp.Utf16 Lean.Data.Lsp.Workspace Lean.Data.Lsp.Ipc Lean.Data.Lsp.CodeActions +// Imports: Lean.Data.Lsp.Basic Lean.Data.Lsp.CancelParams Lean.Data.Lsp.Capabilities Lean.Data.Lsp.Client Lean.Data.Lsp.Communication Lean.Data.Lsp.Diagnostics Lean.Data.Lsp.Extra Lean.Data.Lsp.InitShutdown Lean.Data.Lsp.Internal Lean.Data.Lsp.LanguageFeatures Lean.Data.Lsp.TextSync Lean.Data.Lsp.Utf16 Lean.Data.Lsp.Workspace Lean.Data.Lsp.Ipc Lean.Data.Lsp.CodeActions #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* initialize_Lean_Data_Lsp_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Data_Lsp_CancelParams(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Lsp_Capabilities(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Lsp_Client(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Lsp_Communication(uint8_t builtin, lean_object*); @@ -35,6 +36,9 @@ _G_initialized = true; res = initialize_Lean_Data_Lsp_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Data_Lsp_CancelParams(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Data_Lsp_Capabilities(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Data/Lsp/Basic.c b/stage0/stdlib/Lean/Data/Lsp/Basic.c index 0b3c563a5b67..fbfb4b502ed4 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Basic.c +++ b/stage0/stdlib/Lean/Data/Lsp/Basic.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Data.Lsp.Basic -// Imports: Lean.Data.Json Lean.Data.JsonRpc +// Imports: Lean.Data.Json #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,1358 +13,858 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7; lean_object* l_Lean_JsonNumber_fromNat(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13; static lean_object* l_Lean_Lsp_instToJsonMarkupContent___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_(lean_object*); static lean_object* l_Lean_Lsp_instBEqRange___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CreateFile_instFromJsonOptions; static lean_object* l_Lean_Lsp_instToStringPosition___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonChangeAnnotation; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15; lean_object* l_Lean_Json_getObj_x3f(lean_object*); static lean_object* l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedLocation; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentPositionParams; static lean_object* l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCancelParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1; static lean_object* l_Lean_Lsp_instInhabitedRange___closed__1; static lean_object* l_Lean_Lsp_instFromJsonPartialResultParams___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3; -static lean_object* l_Lean_Lsp_instToJsonCancelParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; static lean_object* l_Lean_Lsp_instHashablePosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentFilter; LEAN_EXPORT lean_object* l_Lean_Lsp_instToStringTextDocumentPositionParams(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1; static lean_object* l_Lean_Lsp_instFromJsonCommand___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCancelParams; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655_(lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__5(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableMarkupContent; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableMarkupKind; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSelector(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__4; -static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1; static lean_object* l_Lean_Lsp_instOrdRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8; uint64_t lean_uint64_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPosition; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12; static lean_object* l_Lean_Lsp_instToJsonTextDocumentItem___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5; static lean_object* l_Lean_Lsp_instToJsonMarkupKind___closed__1; static lean_object* l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonMarkupContent; static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8; -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442____boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10; static lean_object* l_Lean_Lsp_instFromJsonCreateFile___closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_(lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11; LEAN_EXPORT uint8_t l_Lean_Lsp_MarkupKind_ofNat(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressReport; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonTextEditBatch___spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentEdit; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6722_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonPartialResultParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1; static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instOrdPosition; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonLocation___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_ofTextEdit(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5727_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090_(uint8_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentEdit; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressParams; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRenameFile; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15; static lean_object* l_Lean_Lsp_instBEqLocation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1___boxed(lean_object*, lean_object*); uint64_t lean_string_hash(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRange; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8; static lean_object* l_Lean_Lsp_instOrdPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLocation; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_instEmptyCollectionTextEditBatch; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCreateFile___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11; uint8_t lean_string_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3; static lean_object* l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonApplyWorkspaceEditParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqLocation; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLocation; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1; lean_object* l_Lean_Json_getObjVal_x3f(lean_object*, lean_object*); lean_object* l_Lean_Json_getBool_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSnippetString___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_DeleteFile_instToJsonOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10; static lean_object* l_Lean_Lsp_instFromJsonRange___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentIdentifier; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSelector(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextEditBatch(lean_object*); lean_object* l_Lean_Json_getStr_x3f(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentChange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonPosition; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCommand; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentChange(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__5(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__20(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextEdit; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15; static lean_object* l_Lean_Lsp_instFromJsonSnippetString___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983_(lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCommand___closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16; static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1; static lean_object* l_Lean_Lsp_instFromJsonTextEdit___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentPositionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16; LEAN_EXPORT lean_object* l_Lean_Lsp_instLTRange; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12; +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; static lean_object* l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4; static lean_object* l_Lean_Lsp_instHashableMarkupKind___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupKind(uint8_t, uint8_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonTextEditBatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonChangeAnnotation; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_ofTextDocumentEdit(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9; +LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDocumentFilter___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225_(lean_object*); lean_object* l_Lean_Json_setObjVal_x21(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7; static lean_object* l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instEmptyCollection; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instHashablePosition; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19; LEAN_EXPORT lean_object* l_Lean_Lsp_instOrdRange; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCancelParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8; +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3; +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkspaceEdit; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__1; static lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1; static lean_object* l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSnippetString; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentPositionParams; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressBegin; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instCoeTextEditTextEditBatch(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonRange___closed__1; lean_object* lean_array_to_list(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSelector___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonMarkupKind(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonMarkupKind___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToStringPosition(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCreateFile; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12; -LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLocationLink; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkspaceEdit; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166_(lean_object*); static lean_object* l_Lean_Lsp_WorkspaceEdit_instEmptyCollection___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instAppendTextEditBatch; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDeleteFile; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10; static lean_object* l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6633_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonDocumentSelector___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848_(lean_object*); static lean_object* l_Lean_Lsp_instHashableMarkupContent___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPartialResultParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; -LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCancelParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5; LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupContent(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentFilter; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5; +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonPartialResultParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3; static lean_object* l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonMarkupContent; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036_(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonMarkupKind___boxed(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9; static lean_object* l_Lean_Lsp_instOrdLocation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7; +static lean_object* l_Lean_Lsp_instInhabitedLocation___closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instLEPosition; lean_object* l_Array_append___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqPosition; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CreateFile_instToJsonOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11; static lean_object* l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSelector___spec__1(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2; static lean_object* l_Lean_Lsp_instToJsonDeleteFile___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextEdit; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instOrdLocation; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10; -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSnippetString; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2; static lean_object* l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonMarkupKind(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2; -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDeleteFile; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqRange; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2; lean_object* l_Array_append___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__20(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonStaticRegistrationOptions; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instLTPosition; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonStaticRegistrationOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; static lean_object* l_Lean_Lsp_instFromJsonLocation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5; -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6; static lean_object* l_Lean_Lsp_instFromJsonLocationLink___closed__1; static lean_object* l_Lean_Lsp_instInhabitedPosition___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2; static lean_object* l_Lean_Lsp_instFromJsonPosition___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13; static lean_object* l_Lean_Lsp_instInhabitedLocation___closed__1; static lean_object* l_Lean_Lsp_instFromJsonMarkupContent___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7; static lean_object* l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12; static lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DeleteFile_instFromJsonOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqMarkupKind___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6537_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedRange; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6; static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instLERange; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7; -static lean_object* l_Lean_Lsp_instBEqCancelParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8; lean_object* l_Lean_Json_getNat_x3f(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9; static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6; lean_object* l_List_foldl___at_Array_appendList___spec__1___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableRange; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5; static lean_object* l_Lean_Lsp_instAppendTextEditBatch___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3; -LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instBEqPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; -static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3; static lean_object* l_Lean_Lsp_instToJsonRenameFile___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1; lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2; static lean_object* l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(lean_object*); size_t lean_usize_add(size_t, size_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressEnd; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_ofNat___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentItem; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9; static lean_object* l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7; static lean_object* l_Lean_Lsp_instToJsonTextEdit___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedPosition; lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; size_t lean_array_size(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2; static lean_object* l_Lean_Lsp_instToJsonDocumentFilter___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object*, lean_object*); -uint8_t l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_36_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4; static lean_object* l_Lean_Lsp_instToStringPosition___closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqMarkupContent___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4; static lean_object* l_Lean_Lsp_instToJsonChangeAnnotation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCommand; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentItem; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1; lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLocationLink; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12; static lean_object* l_Lean_Lsp_instHashableRange___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRenameFile; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15; +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDeleteFile___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16; lean_object* l_Lean_RBNode_fold___at_Lean_RBMap_mergeBy___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8; static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1___boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCreateFile; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentIdentifier; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____boxed(lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14; -static lean_object* l_Lean_Lsp_instFromJsonCancelParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonMarkupKind___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRange; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; static lean_object* l_Lean_Lsp_instToStringPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1; +LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16; static lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___closed__1; static lean_object* l_Lean_Lsp_instFromJsonRenameFile___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1; +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonDocumentSelector___spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonLocationLink___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextEditBatch(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4; -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6818_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3; -static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Lsp_instInhabitedCancelParams___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instInhabitedCancelParams___closed__2; -return x_1; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28_(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_36_(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28____boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28_(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Lsp_instBEqCancelParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28____boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instBEqCancelParams() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instBEqCancelParams___closed__1; -return x_1; -} -} -LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = lean_array_to_list(x_2); -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -lean_dec(x_1); -x_6 = l_List_foldl___at_Array_appendList___spec__1___rarg(x_2, x_4); -x_1 = x_5; -x_2 = x_6; -goto _start; -} -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("id", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_box(0); -switch (lean_obj_tag(x_1)) { -case 0: -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_ctor_set_tag(x_1, 3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_1); -x_6 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_2); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_2); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); -x_10 = l_Lean_Json_mkObj(x_9); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -lean_dec(x_1); -x_12 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_12, 0, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_2); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_2); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); -x_19 = l_Lean_Json_mkObj(x_18); -return x_19; -} -} -case 1: -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_1); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_ctor_set_tag(x_1, 2); -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_1); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_2); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_2); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_24, x_25); -x_27 = l_Lean_Json_mkObj(x_26); -return x_27; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_28 = lean_ctor_get(x_1, 0); -lean_inc(x_28); -lean_dec(x_1); -x_29 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_2); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_2); -x_34 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_33, x_34); -x_36 = l_Lean_Json_mkObj(x_35); -return x_36; -} -} -default: -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5; -x_38 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_39 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_37, x_38); -x_40 = l_Lean_Json_mkObj(x_39); -return x_40; -} -} -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCancelParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCancelParams() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instToJsonCancelParams___closed__1; -return x_1; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = 0; -return x_2; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lsp", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("CancelParams", 12, 12); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1___boxed), 1, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4; -x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(".", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9; -x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(x_1, x_2); -if (lean_obj_tag(x_3) == 0) -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13; -x_7 = lean_string_append(x_6, x_5); -lean_dec(x_5); -lean_ctor_set(x_3, 0, x_7); -return x_3; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_3, 0); -lean_inc(x_8); -lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13; -x_10 = lean_string_append(x_9, x_8); -lean_dec(x_8); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_3); -if (x_12 == 0) -{ -return x_3; -} -else -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -lean_dec(x_3); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCancelParams___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCancelParams() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instFromJsonCancelParams___closed__1; -return x_1; -} -} +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27(lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Lsp_instInhabitedPosition___closed__1() { _start: { @@ -1384,7 +884,7 @@ x_1 = l_Lean_Lsp_instInhabitedPosition___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -1407,11 +907,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1422,7 +922,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqPosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41____boxed), 2, 0); return x_1; } } @@ -1434,7 +934,7 @@ x_1 = l_Lean_Lsp_instBEqPosition___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -1489,11 +989,11 @@ return x_15; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1504,7 +1004,7 @@ static lean_object* _init_l_Lean_Lsp_instOrdPosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115____boxed), 2, 0); return x_1; } } @@ -1516,7 +1016,7 @@ x_1 = l_Lean_Lsp_instOrdPosition___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint64_t x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; @@ -1530,11 +1030,11 @@ x_8 = lean_uint64_mix_hash(x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -1544,7 +1044,7 @@ static lean_object* _init_l_Lean_Lsp_instHashablePosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180____boxed), 1, 0); return x_1; } } @@ -1556,7 +1056,31 @@ x_1 = l_Lean_Lsp_instHashablePosition___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1() { +LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = lean_array_to_list(x_2); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +lean_dec(x_1); +x_6 = l_List_foldl___at_Array_appendList___spec__1___rarg(x_2, x_4); +x_1 = x_5; +x_2 = x_6; +goto _start; +} +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1() { _start: { lean_object* x_1; @@ -1564,7 +1088,7 @@ x_1 = lean_mk_string_unchecked("line", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2() { _start: { lean_object* x_1; @@ -1572,7 +1096,16 @@ x_1 = lean_mk_string_unchecked("character", 9, 9); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object* x_1) { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -1581,7 +1114,7 @@ lean_inc(x_2); x_3 = l_Lean_JsonNumber_fromNat(x_2); x_4 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_4, 0, x_3); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -1595,7 +1128,7 @@ lean_dec(x_1); x_10 = l_Lean_JsonNumber_fromNat(x_9); x_11 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2; x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); @@ -1608,8 +1141,8 @@ lean_ctor_set(x_15, 1, x_7); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_8); lean_ctor_set(x_16, 1, x_15); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } @@ -1618,7 +1151,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonPosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_), 1, 0); return x_1; } } @@ -1630,7 +1163,7 @@ x_1 = l_Lean_Lsp_instToJsonPosition___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -1639,7 +1172,31 @@ x_4 = l_Lean_Json_getNat_x3f(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1() { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lsp", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3() { _start: { lean_object* x_1; @@ -1647,127 +1204,151 @@ x_1 = lean_mk_string_unchecked("Position", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1777,7 +1358,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -1789,7 +1370,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -1803,8 +1384,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -1814,7 +1395,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -1826,7 +1407,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -1865,20 +1446,30 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} static lean_object* _init_l_Lean_Lsp_instFromJsonPosition___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273_), 1, 0); return x_1; } } @@ -1972,7 +1563,7 @@ x_1 = l_Lean_Lsp_instInhabitedRange___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -1980,7 +1571,7 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_3, x_5); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_3, x_5); if (x_7 == 0) { uint8_t x_8; @@ -1990,16 +1581,16 @@ return x_8; else { uint8_t x_9; -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_4, x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_4, x_6); return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2010,7 +1601,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442____boxed), 2, 0); return x_1; } } @@ -2022,25 +1613,25 @@ x_1 = l_Lean_Lsp_instBEqRange___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint64_t x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; x_2 = lean_ctor_get(x_1, 0); x_3 = lean_ctor_get(x_1, 1); x_4 = 0; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_2); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_2); x_6 = lean_uint64_mix_hash(x_4, x_5); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_3); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_3); x_8 = lean_uint64_mix_hash(x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -2050,7 +1641,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516____boxed), 1, 0); return x_1; } } @@ -2062,7 +1653,7 @@ x_1 = l_Lean_Lsp_instHashableRange___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1() { _start: { lean_object* x_1; @@ -2070,7 +1661,7 @@ x_1 = lean_mk_string_unchecked("start", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2() { _start: { lean_object* x_1; @@ -2078,14 +1669,14 @@ x_1 = lean_mk_string_unchecked("end", 3, 3); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_2); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -2096,8 +1687,8 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -2110,8 +1701,8 @@ lean_ctor_set(x_13, 1, x_6); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -2120,7 +1711,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_), 1, 0); return x_1; } } @@ -2132,16 +1723,16 @@ x_1 = l_Lean_Lsp_instToJsonRange___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1() { _start: { lean_object* x_1; @@ -2149,127 +1740,127 @@ x_1 = lean_mk_string_unchecked("Range", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2279,7 +1870,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -2291,7 +1882,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -2305,8 +1896,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2316,7 +1907,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -2328,7 +1919,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -2367,11 +1958,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -2380,7 +1971,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_), 1, 0); return x_1; } } @@ -2392,7 +1983,7 @@ x_1 = l_Lean_Lsp_instFromJsonRange___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; uint8_t x_8; lean_object* x_9; @@ -2400,8 +1991,8 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_3, x_5); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_4, x_6); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_3, x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_4, x_6); x_9 = lean_box(x_8); if (lean_obj_tag(x_9) == 1) { @@ -2436,11 +2027,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2451,7 +2042,7 @@ static lean_object* _init_l_Lean_Lsp_instOrdRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715____boxed), 2, 0); return x_1; } } @@ -2482,8 +2073,16 @@ return x_1; static lean_object* _init_l_Lean_Lsp_instInhabitedLocation___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instInhabitedLocation___closed__2() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Lsp_instInhabitedCancelParams___closed__1; +x_1 = l_Lean_Lsp_instInhabitedLocation___closed__1; x_2 = l_Lean_Lsp_instInhabitedRange___closed__1; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2495,11 +2094,11 @@ static lean_object* _init_l_Lean_Lsp_instInhabitedLocation() { _start: { lean_object* x_1; -x_1 = l_Lean_Lsp_instInhabitedLocation___closed__1; +x_1 = l_Lean_Lsp_instInhabitedLocation___closed__2; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -2517,16 +2116,16 @@ return x_8; else { uint8_t x_9; -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_6); return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2537,7 +2136,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqLocation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821____boxed), 2, 0); return x_1; } } @@ -2549,7 +2148,7 @@ x_1 = l_Lean_Lsp_instBEqLocation___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1() { _start: { lean_object* x_1; @@ -2557,7 +2156,7 @@ x_1 = lean_mk_string_unchecked("uri", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2() { _start: { lean_object* x_1; @@ -2565,7 +2164,7 @@ x_1 = lean_mk_string_unchecked("range", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object* x_1) { _start: { uint8_t x_2; @@ -2577,15 +2176,15 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_3); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_ctor_set(x_1, 1, x_5); lean_ctor_set(x_1, 0, x_6); x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -2598,8 +2197,8 @@ lean_ctor_set(x_13, 1, x_7); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -2613,7 +2212,7 @@ lean_inc(x_18); lean_dec(x_1); x_20 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_20, 0, x_18); -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); @@ -2621,8 +2220,8 @@ x_23 = lean_box(0); x_24 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_19); -x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_19); +x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_25); @@ -2635,8 +2234,8 @@ lean_ctor_set(x_29, 1, x_23); x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_30, x_31); +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_30, x_31); x_33 = l_Lean_Json_mkObj(x_32); return x_33; } @@ -2646,7 +2245,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonLocation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_), 1, 0); return x_1; } } @@ -2658,7 +2257,7 @@ x_1 = l_Lean_Lsp_instToJsonLocation___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -2667,16 +2266,16 @@ x_4 = l_Lean_Json_getStr_x3f(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1() { _start: { lean_object* x_1; @@ -2684,127 +2283,127 @@ x_1 = lean_mk_string_unchecked("Location", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2814,7 +2413,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -2826,7 +2425,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -2840,8 +2439,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2851,7 +2450,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -2863,7 +2462,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -2902,20 +2501,20 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } @@ -2924,7 +2523,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonLocation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_), 1, 0); return x_1; } } @@ -2936,7 +2535,7 @@ x_1 = l_Lean_Lsp_instFromJsonLocation___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -2948,7 +2547,7 @@ x_7 = lean_string_dec_lt(x_3, x_5); if (x_7 == 0) { uint8_t x_8; uint8_t x_9; -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_4, x_6); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_4, x_6); x_9 = lean_string_dec_eq(x_3, x_5); if (x_9 == 0) { @@ -2981,11 +2580,11 @@ return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2996,7 +2595,7 @@ static lean_object* _init_l_Lean_Lsp_instOrdLocation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053____boxed), 2, 0); return x_1; } } @@ -3008,7 +2607,7 @@ x_1 = l_Lean_Lsp_instOrdLocation___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -3024,7 +2623,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -3036,7 +2635,7 @@ return x_8; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1() { _start: { lean_object* x_1; @@ -3044,7 +2643,7 @@ x_1 = lean_mk_string_unchecked("originSelectionRange", 20, 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2() { _start: { lean_object* x_1; @@ -3052,7 +2651,7 @@ x_1 = lean_mk_string_unchecked("targetUri", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3() { _start: { lean_object* x_1; @@ -3060,7 +2659,7 @@ x_1 = lean_mk_string_unchecked("targetRange", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4() { _start: { lean_object* x_1; @@ -3068,19 +2667,19 @@ x_1 = lean_mk_string_unchecked("targetSelectionRange", 20, 20); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_3, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1; +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_5); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -3090,8 +2689,8 @@ lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = lean_ctor_get(x_1, 2); lean_inc(x_11); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_11); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -3101,8 +2700,8 @@ lean_ctor_set(x_15, 1, x_9); x_16 = lean_ctor_get(x_1, 3); lean_inc(x_16); lean_dec(x_1); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_16); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_16); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4; x_19 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_17); @@ -3121,8 +2720,8 @@ lean_ctor_set(x_23, 1, x_22); x_24 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_24, 0, x_4); lean_ctor_set(x_24, 1, x_23); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_24, x_25); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_24, x_25); x_27 = l_Lean_Json_mkObj(x_26); return x_27; } @@ -3131,7 +2730,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonLocationLink___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_), 1, 0); return x_1; } } @@ -3143,7 +2742,7 @@ x_1 = l_Lean_Lsp_instToJsonLocationLink___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -3153,7 +2752,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -3162,13 +2761,13 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -3219,7 +2818,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -3323,7 +2922,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1() { _start: { lean_object* x_1; @@ -3331,39 +2930,39 @@ x_1 = lean_mk_string_unchecked("LocationLink", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5() { _start: { lean_object* x_1; @@ -3371,177 +2970,177 @@ x_1 = lean_mk_string_unchecked("originSelectionRange\?", 21, 21); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3551,7 +3150,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -3563,7 +3162,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -3577,9 +3176,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -3590,7 +3189,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -3602,7 +3201,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -3616,9 +3215,9 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3; +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -3630,7 +3229,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -3642,7 +3241,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -3656,8 +3255,8 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -3669,7 +3268,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21; +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -3681,7 +3280,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -3726,11 +3325,11 @@ return x_50; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -3739,7 +3338,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonLocationLink___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225_), 1, 0); return x_1; } } @@ -3751,7 +3350,7 @@ x_1 = l_Lean_Lsp_instFromJsonLocationLink___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -3775,7 +3374,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -3795,7 +3394,7 @@ lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; le x_5 = lean_ctor_get(x_2, 0); x_6 = lean_array_size(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_6, x_7, x_5); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_8); x_9 = lean_alloc_ctor(0, 2, 0); @@ -3815,7 +3414,7 @@ lean_inc(x_12); lean_dec(x_2); x_13 = lean_array_size(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_13, x_14, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_13, x_14, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = lean_alloc_ctor(0, 2, 0); @@ -3830,7 +3429,7 @@ return x_19; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1() { _start: { lean_object* x_1; @@ -3838,7 +3437,7 @@ x_1 = lean_mk_string_unchecked("title", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2() { _start: { lean_object* x_1; @@ -3846,7 +3445,7 @@ x_1 = lean_mk_string_unchecked("command", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3() { _start: { lean_object* x_1; @@ -3854,7 +3453,7 @@ x_1 = lean_mk_string_unchecked("arguments", 9, 9); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -3867,7 +3466,7 @@ lean_inc(x_4); lean_dec(x_1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_2); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1; x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -3877,15 +3476,15 @@ lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); x_10 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_10, 0, x_3); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_8); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__1(x_14, x_4); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3; +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__1(x_14, x_4); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_8); @@ -3895,13 +3494,13 @@ lean_ctor_set(x_17, 1, x_16); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_9); lean_ctor_set(x_18, 1, x_17); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_20 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_18, x_19); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_20 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_18, x_19); x_21 = l_Lean_Json_mkObj(x_20); return x_21; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -3909,7 +3508,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_4, x_5, x_3); return x_6; } } @@ -3917,7 +3516,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCommand___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_), 1, 0); return x_1; } } @@ -3929,7 +3528,16 @@ x_1 = l_Lean_Lsp_instToJsonCommand___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Lean_Json_getObjValD(x_1, x_2); +x_4 = l_Lean_Json_getStr_x3f(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -3956,7 +3564,7 @@ goto _start; } } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1() { _start: { lean_object* x_1; @@ -3964,7 +3572,7 @@ x_1 = lean_mk_string_unchecked("expected JSON array, got '", 26, 26); return x_1; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2() { _start: { lean_object* x_1; @@ -3972,7 +3580,7 @@ x_1 = lean_mk_string_unchecked("'", 1, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -3981,7 +3589,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -3989,10 +3597,10 @@ case 1: lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_5 = lean_unsigned_to_nat(80u); x_6 = l_Lean_Json_pretty(x_3, x_5); -x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_8 = lean_string_append(x_7, x_6); lean_dec(x_6); -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_10 = lean_string_append(x_8, x_9); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); @@ -4008,7 +3616,7 @@ lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; uint8_t x_17; x_13 = lean_ctor_get(x_3, 0); x_14 = lean_array_size(x_13); x_15 = 0; -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_14, x_15, x_13); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_14, x_15, x_13); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { @@ -4040,7 +3648,7 @@ lean_inc(x_21); lean_dec(x_3); x_22 = lean_array_size(x_21); x_23 = 0; -x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_22, x_23, x_21); +x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_22, x_23, x_21); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); if (lean_is_exclusive(x_24)) { @@ -4073,10 +3681,10 @@ if (x_31 == 0) lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_32 = lean_ctor_get(x_3, 0); lean_dec(x_32); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_30); lean_dec(x_30); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_36); @@ -4086,10 +3694,10 @@ else { lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_dec(x_3); -x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_38 = lean_string_append(x_37, x_30); lean_dec(x_30); -x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_40 = lean_string_append(x_38, x_39); x_41 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_41, 0, x_40); @@ -4099,7 +3707,7 @@ return x_41; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1() { _start: { lean_object* x_1; @@ -4107,121 +3715,121 @@ x_1 = lean_mk_string_unchecked("Command", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13() { _start: { lean_object* x_1; @@ -4229,54 +3837,54 @@ x_1 = lean_mk_string_unchecked("arguments\?", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4286,7 +3894,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -4298,7 +3906,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -4312,9 +3920,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -4325,7 +3933,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -4337,7 +3945,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -4351,8 +3959,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -4363,7 +3971,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -4375,7 +3983,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -4417,7 +4025,16 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -4425,15 +4042,15 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4442,7 +4059,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCommand___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_), 1, 0); return x_1; } } @@ -4454,7 +4071,7 @@ x_1 = l_Lean_Lsp_instFromJsonCommand___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1() { _start: { lean_object* x_1; @@ -4462,13 +4079,13 @@ x_1 = lean_mk_string_unchecked("value", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -4479,8 +4096,8 @@ lean_ctor_set(x_6, 1, x_5); x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -4489,7 +4106,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSnippetString___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655_), 1, 0); return x_1; } } @@ -4501,7 +4118,7 @@ x_1 = l_Lean_Lsp_instToJsonSnippetString___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1() { _start: { lean_object* x_1; @@ -4509,85 +4126,85 @@ x_1 = lean_mk_string_unchecked("SnippetString", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4596,7 +4213,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -4608,7 +4225,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -4641,7 +4258,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSnippetString___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_), 1, 0); return x_1; } } @@ -4653,7 +4270,7 @@ x_1 = l_Lean_Lsp_instFromJsonSnippetString___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4669,7 +4286,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -4681,7 +4298,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4728,7 +4345,7 @@ return x_12; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1() { _start: { lean_object* x_1; @@ -4736,7 +4353,7 @@ x_1 = lean_mk_string_unchecked("newText", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2() { _start: { lean_object* x_1; @@ -4744,7 +4361,7 @@ x_1 = lean_mk_string_unchecked("leanExtSnippet", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3() { _start: { lean_object* x_1; @@ -4752,14 +4369,14 @@ x_1 = lean_mk_string_unchecked("annotationId", 12, 12); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_2); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -4771,7 +4388,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); x_9 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_9, 0, x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -4780,13 +4397,13 @@ lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); x_13 = lean_ctor_get(x_1, 2); lean_inc(x_13); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__1(x_14, x_13); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2; +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__1(x_14, x_13); x_16 = lean_ctor_get(x_1, 3); lean_inc(x_16); lean_dec(x_1); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_17, x_16); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_17, x_16); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_6); @@ -4799,8 +4416,8 @@ lean_ctor_set(x_21, 1, x_20); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_7); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_22, x_23); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_22, x_23); x_25 = l_Lean_Json_mkObj(x_24); return x_25; } @@ -4809,7 +4426,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800_), 1, 0); return x_1; } } @@ -4821,7 +4438,7 @@ x_1 = l_Lean_Lsp_instToJsonTextEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -4830,13 +4447,13 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -4887,7 +4504,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -4991,7 +4608,7 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -5000,7 +4617,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -5161,7 +4778,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1() { _start: { lean_object* x_1; @@ -5169,100 +4786,100 @@ x_1 = lean_mk_string_unchecked("TextEdit", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11() { _start: { lean_object* x_1; @@ -5270,48 +4887,48 @@ x_1 = lean_mk_string_unchecked("leanExtSnippet\?", 15, 15); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16() { _start: { lean_object* x_1; @@ -5319,54 +4936,54 @@ x_1 = lean_mk_string_unchecked("annotationId\?", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5376,7 +4993,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -5388,7 +5005,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -5402,9 +5019,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -5415,7 +5032,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -5427,7 +5044,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -5441,9 +5058,9 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2; +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -5455,7 +5072,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -5467,7 +5084,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -5481,8 +5098,8 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -5494,7 +5111,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20; +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -5506,7 +5123,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -5551,20 +5168,20 @@ return x_50; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } @@ -5573,7 +5190,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860_), 1, 0); return x_1; } } @@ -5603,7 +5220,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_6 = lean_array_uget(x_3, x_2); x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_3, x_2, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045_(x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860_(x_6); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -5649,10 +5266,10 @@ case 0: lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_2 = lean_unsigned_to_nat(80u); x_3 = l_Lean_Json_pretty(x_1, x_2); -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_5 = lean_string_append(x_4, x_3); lean_dec(x_3); -x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_7 = lean_string_append(x_5, x_6); x_8 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_8, 0, x_7); @@ -5663,10 +5280,10 @@ case 1: lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_9 = lean_unsigned_to_nat(80u); x_10 = l_Lean_Json_pretty(x_1, x_9); -x_11 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_11 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_12 = lean_string_append(x_11, x_10); lean_dec(x_10); -x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_14 = lean_string_append(x_12, x_13); x_15 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_15, 0, x_14); @@ -5695,10 +5312,10 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_23 = lean_ctor_get(x_1, 0); lean_dec(x_23); -x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_25 = lean_string_append(x_24, x_21); lean_dec(x_21); -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_27 = lean_string_append(x_25, x_26); lean_ctor_set_tag(x_1, 0); lean_ctor_set(x_1, 0, x_27); @@ -5708,10 +5325,10 @@ else { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_1); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_21); lean_dec(x_21); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); x_32 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_32, 0, x_31); @@ -5748,7 +5365,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -5786,7 +5403,7 @@ static lean_object* _init_l_Lean_Lsp_instEmptyCollectionTextEditBatch() { _start: { lean_object* x_1; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; return x_1; } } @@ -5818,13 +5435,13 @@ x_4 = lean_array_mk(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -5835,8 +5452,8 @@ lean_ctor_set(x_6, 1, x_5); x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -5845,7 +5462,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_), 1, 0); return x_1; } } @@ -5857,7 +5474,7 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1() { _start: { lean_object* x_1; @@ -5865,64 +5482,64 @@ x_1 = lean_mk_string_unchecked("TextDocumentIdentifier", 22, 22); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5931,7 +5548,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -5943,7 +5560,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -5976,7 +5593,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166_), 1, 0); return x_1; } } @@ -5988,7 +5605,7 @@ x_1 = l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -6039,7 +5656,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1() { _start: { lean_object* x_1; @@ -6047,7 +5664,7 @@ x_1 = lean_mk_string_unchecked("version", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(lean_object* x_1) { _start: { uint8_t x_2; @@ -6059,23 +5676,23 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_3); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_ctor_set(x_1, 1, x_5); lean_ctor_set(x_1, 0, x_6); x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_9, x_4); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_8); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -6089,7 +5706,7 @@ lean_inc(x_16); lean_dec(x_1); x_18 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_18, 0, x_16); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_20 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_18); @@ -6097,16 +5714,16 @@ x_21 = lean_box(0); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_23, x_17); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_21); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_22); lean_ctor_set(x_26, 1, x_25); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } @@ -6116,7 +5733,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_), 1, 0); return x_1; } } @@ -6128,7 +5745,7 @@ x_1 = l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6137,7 +5754,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -6298,7 +5915,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1() { _start: { lean_object* x_1; @@ -6306,59 +5923,59 @@ x_1 = lean_mk_string_unchecked("VersionedTextDocumentIdentifier", 31, 31); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7() { _start: { lean_object* x_1; @@ -6366,54 +5983,54 @@ x_1 = lean_mk_string_unchecked("version\?", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -6423,7 +6040,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -6435,7 +6052,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -6449,8 +6066,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -6460,7 +6077,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -6472,7 +6089,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -6511,11 +6128,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -6524,7 +6141,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297_), 1, 0); return x_1; } } @@ -6536,7 +6153,7 @@ x_1 = l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1() { _start: { lean_object* x_1; @@ -6544,7 +6161,7 @@ x_1 = lean_mk_string_unchecked("textDocument", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2() { _start: { lean_object* x_1; @@ -6552,14 +6169,14 @@ x_1 = lean_mk_string_unchecked("edits", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(x_2); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -6575,7 +6192,7 @@ x_10 = 0; x_11 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(x_9, x_10, x_8); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -6588,8 +6205,8 @@ lean_ctor_set(x_16, 1, x_6); x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_7); lean_ctor_set(x_17, 1, x_16); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -6598,7 +6215,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_), 1, 0); return x_1; } } @@ -6610,16 +6227,16 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297_(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6630,10 +6247,10 @@ case 0: lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_4 = lean_unsigned_to_nat(80u); x_5 = l_Lean_Json_pretty(x_3, x_4); -x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_8 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_10, 0, x_9); @@ -6644,10 +6261,10 @@ case 1: lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_11 = lean_unsigned_to_nat(80u); x_12 = l_Lean_Json_pretty(x_3, x_11); -x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_14 = lean_string_append(x_13, x_12); lean_dec(x_12); -x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -6676,10 +6293,10 @@ if (x_24 == 0) lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; x_25 = lean_ctor_get(x_3, 0); lean_dec(x_25); -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_27 = lean_string_append(x_26, x_23); lean_dec(x_23); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_29 = lean_string_append(x_27, x_28); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_29); @@ -6689,10 +6306,10 @@ else { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_dec(x_3); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_31 = lean_string_append(x_30, x_23); lean_dec(x_23); -x_32 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_32 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_33 = lean_string_append(x_31, x_32); x_34 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_34, 0, x_33); @@ -6702,7 +6319,7 @@ return x_34; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1() { _start: { lean_object* x_1; @@ -6710,127 +6327,127 @@ x_1 = lean_mk_string_unchecked("TextDocumentEdit", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -6840,7 +6457,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -6852,7 +6469,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -6866,8 +6483,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -6877,7 +6494,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -6889,7 +6506,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -6928,20 +6545,20 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } @@ -6950,7 +6567,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_), 1, 0); return x_1; } } @@ -6962,7 +6579,7 @@ x_1 = l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1() { _start: { lean_object* x_1; @@ -6970,7 +6587,7 @@ x_1 = lean_mk_string_unchecked("label", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2() { _start: { lean_object* x_1; @@ -6978,7 +6595,7 @@ x_1 = lean_mk_string_unchecked("needsConfirmation", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3() { _start: { lean_object* x_1; @@ -6986,7 +6603,7 @@ x_1 = lean_mk_string_unchecked("description", 11, 11); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -6998,7 +6615,7 @@ lean_inc(x_4); lean_dec(x_1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_2); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -7008,15 +6625,15 @@ lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); x_10 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_10, 0, x_3); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2; +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_8); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_14, x_4); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3; +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_14, x_4); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_8); @@ -7026,8 +6643,8 @@ lean_ctor_set(x_17, 1, x_16); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_9); lean_ctor_set(x_18, 1, x_17); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_20 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_18, x_19); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_20 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_18, x_19); x_21 = l_Lean_Json_mkObj(x_20); return x_21; } @@ -7036,7 +6653,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonChangeAnnotation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_), 1, 0); return x_1; } } @@ -7048,7 +6665,7 @@ x_1 = l_Lean_Lsp_instToJsonChangeAnnotation___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -7058,7 +6675,7 @@ lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1() { _start: { lean_object* x_1; @@ -7066,121 +6683,121 @@ x_1 = lean_mk_string_unchecked("ChangeAnnotation", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13() { _start: { lean_object* x_1; @@ -7188,54 +6805,54 @@ x_1 = lean_mk_string_unchecked("description\?", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7245,7 +6862,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7257,7 +6874,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7271,9 +6888,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7284,7 +6901,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7296,7 +6913,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -7310,8 +6927,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -7322,7 +6939,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -7334,7 +6951,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -7380,11 +6997,11 @@ return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -7393,7 +7010,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_), 1, 0); return x_1; } } @@ -7405,7 +7022,7 @@ x_1 = l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1() { _start: { lean_object* x_1; @@ -7413,7 +7030,7 @@ x_1 = lean_mk_string_unchecked("overwrite", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2() { _start: { lean_object* x_1; @@ -7421,7 +7038,7 @@ x_1 = lean_mk_string_unchecked("ignoreIfExists", 14, 14); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839_(lean_object* x_1) { _start: { uint8_t x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -7429,7 +7046,7 @@ x_2 = lean_ctor_get_uint8(x_1, 0); x_3 = lean_ctor_get_uint8(x_1, 1); x_4 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_4, 0, x_2); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1; +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -7439,7 +7056,7 @@ lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); x_9 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_9, 0, x_3); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -7452,17 +7069,17 @@ lean_ctor_set(x_13, 1, x_7); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839_(x_1); lean_dec(x_1); return x_2; } @@ -7471,7 +7088,7 @@ static lean_object* _init_l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____boxed), 1, 0); return x_1; } } @@ -7483,7 +7100,7 @@ x_1 = l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1() { _start: { lean_object* x_1; @@ -7491,7 +7108,7 @@ x_1 = lean_mk_string_unchecked("CreateFile", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2() { _start: { lean_object* x_1; @@ -7499,128 +7116,128 @@ x_1 = lean_mk_string_unchecked("Options", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1; -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1; +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7630,7 +7247,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7642,7 +7259,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7656,8 +7273,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7667,7 +7284,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7679,7 +7296,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -7730,7 +7347,7 @@ static lean_object* _init_l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_), 1, 0); return x_1; } } @@ -7742,7 +7359,7 @@ x_1 = l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1() { _start: { lean_object* x_1; @@ -7750,7 +7367,7 @@ x_1 = lean_mk_string_unchecked("recursive", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2() { _start: { lean_object* x_1; @@ -7758,7 +7375,7 @@ x_1 = lean_mk_string_unchecked("ignoreIfNotExists", 17, 17); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023_(lean_object* x_1) { _start: { uint8_t x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -7766,7 +7383,7 @@ x_2 = lean_ctor_get_uint8(x_1, 0); x_3 = lean_ctor_get_uint8(x_1, 1); x_4 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_4, 0, x_2); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1; +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -7776,7 +7393,7 @@ lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); x_9 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_9, 0, x_3); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -7789,17 +7406,17 @@ lean_ctor_set(x_13, 1, x_7); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023_(x_1); lean_dec(x_1); return x_2; } @@ -7808,7 +7425,7 @@ static lean_object* _init_l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____boxed), 1, 0); return x_1; } } @@ -7820,7 +7437,7 @@ x_1 = l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1() { _start: { lean_object* x_1; @@ -7828,128 +7445,128 @@ x_1 = lean_mk_string_unchecked("DeleteFile", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1; +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7959,7 +7576,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7971,7 +7588,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7985,8 +7602,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7996,7 +7613,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8008,7 +7625,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8059,7 +7676,7 @@ static lean_object* _init_l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_), 1, 0); return x_1; } } @@ -8071,7 +7688,7 @@ x_1 = l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -8085,7 +7702,7 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_2, 0); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -8097,7 +7714,7 @@ return x_8; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1() { _start: { lean_object* x_1; @@ -8105,7 +7722,7 @@ x_1 = lean_mk_string_unchecked("options", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -8118,7 +7735,7 @@ lean_inc(x_4); lean_dec(x_1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_2); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -8126,11 +7743,11 @@ x_8 = lean_box(0); x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; -x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(x_10, x_3); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; +x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(x_10, x_3); lean_dec(x_3); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_12, x_4); +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_12, x_4); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_8); @@ -8140,17 +7757,17 @@ lean_ctor_set(x_15, 1, x_14); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_9); lean_ctor_set(x_16, 1, x_15); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -8159,7 +7776,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCreateFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_), 1, 0); return x_1; } } @@ -8171,7 +7788,7 @@ x_1 = l_Lean_Lsp_instToJsonCreateFile___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -8180,13 +7797,13 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -8237,7 +7854,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -8341,59 +7958,59 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6() { _start: { lean_object* x_1; @@ -8401,74 +8018,74 @@ x_1 = lean_mk_string_unchecked("options\?", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -8478,7 +8095,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8490,7 +8107,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8504,9 +8121,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -8517,7 +8134,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8529,7 +8146,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8543,8 +8160,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -8555,7 +8172,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -8567,7 +8184,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -8609,11 +8226,11 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -8622,7 +8239,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCreateFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_), 1, 0); return x_1; } } @@ -8634,7 +8251,7 @@ x_1 = l_Lean_Lsp_instFromJsonCreateFile___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1() { _start: { lean_object* x_1; @@ -8642,7 +8259,7 @@ x_1 = lean_mk_string_unchecked("oldUri", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2() { _start: { lean_object* x_1; @@ -8650,7 +8267,7 @@ x_1 = lean_mk_string_unchecked("newUri", 6, 6); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; @@ -8665,7 +8282,7 @@ lean_inc(x_5); lean_dec(x_1); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_2); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -8675,18 +8292,18 @@ lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_11, 0, x_3); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_9); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; -x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____spec__1(x_15, x_4); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; +x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____spec__1(x_15, x_4); lean_dec(x_4); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_17, x_5); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_17, x_5); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_9); @@ -8699,8 +8316,8 @@ lean_ctor_set(x_21, 1, x_20); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_10); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_22, x_23); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_22, x_23); x_25 = l_Lean_Json_mkObj(x_24); return x_25; } @@ -8709,7 +8326,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonRenameFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_), 1, 0); return x_1; } } @@ -8721,7 +8338,7 @@ x_1 = l_Lean_Lsp_instToJsonRenameFile___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1() { _start: { lean_object* x_1; @@ -8729,167 +8346,167 @@ x_1 = lean_mk_string_unchecked("RenameFile", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -8899,7 +8516,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8911,7 +8528,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8925,9 +8542,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -8938,7 +8555,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8950,7 +8567,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8964,9 +8581,9 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -8978,7 +8595,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -8990,7 +8607,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -9004,8 +8621,8 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -9017,7 +8634,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16; +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -9029,7 +8646,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -9078,7 +8695,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonRenameFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_), 1, 0); return x_1; } } @@ -9090,7 +8707,7 @@ x_1 = l_Lean_Lsp_instFromJsonRenameFile___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -9104,7 +8721,7 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_2, 0); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -9116,7 +8733,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -9129,7 +8746,7 @@ lean_inc(x_4); lean_dec(x_1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_2); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -9137,11 +8754,11 @@ x_8 = lean_box(0); x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; -x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1(x_10, x_3); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; +x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1(x_10, x_3); lean_dec(x_3); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_12, x_4); +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_12, x_4); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_8); @@ -9151,17 +8768,17 @@ lean_ctor_set(x_15, 1, x_14); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_9); lean_ctor_set(x_16, 1, x_15); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -9170,7 +8787,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDeleteFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_), 1, 0); return x_1; } } @@ -9182,7 +8799,7 @@ x_1 = l_Lean_Lsp_instToJsonDeleteFile___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -9191,13 +8808,13 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -9248,7 +8865,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -9352,105 +8969,105 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -9460,7 +9077,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9472,7 +9089,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9486,9 +9103,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -9499,7 +9116,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -9511,7 +9128,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -9525,8 +9142,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -9537,7 +9154,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -9549,7 +9166,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -9591,11 +9208,11 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -9604,7 +9221,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDeleteFile___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_), 1, 0); return x_1; } } @@ -9688,7 +9305,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_(x_2); x_4 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_5 = l_Lean_Lsp_instToJsonDocumentChange___closed__2; x_6 = l_Lean_Json_setObjVal_x21(x_3, x_4, x_5); @@ -9700,7 +9317,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_ctor_get(x_1, 0); lean_inc(x_7); lean_dec(x_1); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_(x_7); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_7); x_9 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_10 = l_Lean_Lsp_instToJsonDocumentChange___closed__5; x_11 = l_Lean_Json_setObjVal_x21(x_8, x_9, x_10); @@ -9712,7 +9329,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean x_12 = lean_ctor_get(x_1, 0); lean_inc(x_12); lean_dec(x_1); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_(x_12); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_(x_12); x_14 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_15 = l_Lean_Lsp_instToJsonDocumentChange___closed__7; x_16 = l_Lean_Json_setObjVal_x21(x_13, x_14, x_15); @@ -9724,7 +9341,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_1, 0); lean_inc(x_17); lean_dec(x_1); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_(x_17); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_(x_17); return x_18; } } @@ -9784,7 +9401,7 @@ else { lean_object* x_27; lean_inc(x_1); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(x_1); +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(x_1); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; @@ -9828,7 +9445,7 @@ else lean_object* x_33; lean_dec(x_19); lean_inc(x_1); -x_33 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(x_1); +x_33 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(x_1); if (lean_obj_tag(x_33) == 0) { lean_object* x_34; @@ -9872,7 +9489,7 @@ else lean_object* x_39; lean_dec(x_19); lean_inc(x_1); -x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_1); +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_1); if (lean_obj_tag(x_39) == 0) { lean_object* x_40; @@ -9941,7 +9558,7 @@ else { lean_object* x_53; lean_inc(x_1); -x_53 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(x_1); +x_53 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(x_1); if (lean_obj_tag(x_53) == 0) { lean_object* x_54; @@ -9980,7 +9597,7 @@ else lean_object* x_59; lean_dec(x_45); lean_inc(x_1); -x_59 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(x_1); +x_59 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(x_1); if (lean_obj_tag(x_59) == 0) { lean_object* x_60; @@ -10019,7 +9636,7 @@ else lean_object* x_65; lean_dec(x_45); lean_inc(x_1); -x_65 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_1); +x_65 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_1); if (lean_obj_tag(x_65) == 0) { lean_object* x_66; @@ -10067,7 +9684,7 @@ goto block_13; { lean_object* x_3; lean_dec(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_(x_1); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_(x_1); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10116,7 +9733,7 @@ return x_12; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10135,13 +9752,13 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x x_4 = lean_ctor_get(x_1, 0); x_5 = lean_ctor_get(x_1, 2); x_6 = lean_ctor_get(x_1, 3); -x_7 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_4); +x_7 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_4); x_8 = lean_array_size(x_5); x_9 = 0; x_10 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(x_8, x_9, x_5); x_11 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_6); +x_12 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_6); lean_ctor_set(x_1, 3, x_12); lean_ctor_set(x_1, 2, x_11); lean_ctor_set(x_1, 0, x_7); @@ -10160,13 +9777,13 @@ lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_dec(x_1); -x_18 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_14); +x_18 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_14); x_19 = lean_array_size(x_16); x_20 = 0; x_21 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(x_19, x_20, x_16); x_22 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_22, 0, x_21); -x_23 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_17); +x_23 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_17); x_24 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_24, 0, x_18); lean_ctor_set(x_24, 1, x_15); @@ -10178,7 +9795,7 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -10196,7 +9813,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_5 = lean_ctor_get(x_2, 0); -x_6 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_5); +x_6 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_5); lean_ctor_set_tag(x_2, 5); lean_ctor_set(x_2, 0, x_6); x_7 = lean_alloc_ctor(0, 2, 0); @@ -10214,7 +9831,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean x_10 = lean_ctor_get(x_2, 0); lean_inc(x_10); lean_dec(x_2); -x_11 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__2(x_10); +x_11 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__2(x_10); x_12 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = lean_alloc_ctor(0, 2, 0); @@ -10229,7 +9846,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -10253,7 +9870,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean x_10 = lean_ctor_get(x_5, 0); lean_inc(x_10); lean_dec(x_5); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399_(x_10); +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214_(x_10); x_12 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_13 = l_Lean_Lsp_instToJsonDocumentChange___closed__2; x_14 = l_Lean_Json_setObjVal_x21(x_11, x_12, x_13); @@ -10268,7 +9885,7 @@ lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean x_17 = lean_ctor_get(x_5, 0); lean_inc(x_17); lean_dec(x_5); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630_(x_17); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_17); x_19 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_20 = l_Lean_Lsp_instToJsonDocumentChange___closed__5; x_21 = l_Lean_Json_setObjVal_x21(x_18, x_19, x_20); @@ -10283,7 +9900,7 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean x_24 = lean_ctor_get(x_5, 0); lean_inc(x_24); lean_dec(x_5); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3906_(x_24); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3721_(x_24); x_26 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; x_27 = l_Lean_Lsp_instToJsonDocumentChange___closed__7; x_28 = l_Lean_Json_setObjVal_x21(x_25, x_26, x_27); @@ -10298,7 +9915,7 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; x_31 = lean_ctor_get(x_5, 0); lean_inc(x_31); lean_dec(x_5); -x_32 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608_(x_31); +x_32 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423_(x_31); x_33 = lean_array_uset(x_7, x_2, x_32); x_2 = x_9; x_3 = x_33; @@ -10308,7 +9925,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -10328,7 +9945,7 @@ lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; le x_5 = lean_ctor_get(x_2, 0); x_6 = lean_array_size(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(x_6, x_7, x_5); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_8); x_9 = lean_alloc_ctor(0, 2, 0); @@ -10348,7 +9965,7 @@ lean_inc(x_12); lean_dec(x_2); x_13 = lean_array_size(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(x_13, x_14, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(x_13, x_14, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = lean_alloc_ctor(0, 2, 0); @@ -10363,7 +9980,7 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10382,9 +9999,9 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_1, 0); x_5 = lean_ctor_get(x_1, 2); x_6 = lean_ctor_get(x_1, 3); -x_7 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_4); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_(x_5); -x_9 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_6); +x_7 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_4); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_(x_5); +x_9 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_6); lean_ctor_set(x_1, 3, x_9); lean_ctor_set(x_1, 2, x_8); lean_ctor_set(x_1, 0, x_7); @@ -10403,9 +10020,9 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_dec(x_1); -x_15 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_11); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797_(x_13); -x_17 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_14); +x_15 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_11); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612_(x_13); +x_17 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_14); x_18 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_18, 0, x_15); lean_ctor_set(x_18, 1, x_12); @@ -10417,7 +10034,7 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__5(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -10435,7 +10052,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_5 = lean_ctor_get(x_2, 0); -x_6 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_5); +x_6 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_5); lean_ctor_set_tag(x_2, 5); lean_ctor_set(x_2, 0, x_6); x_7 = lean_alloc_ctor(0, 2, 0); @@ -10453,7 +10070,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean x_10 = lean_ctor_get(x_2, 0); lean_inc(x_10); lean_dec(x_2); -x_11 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__6(x_10); +x_11 = l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__6(x_10); x_12 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = lean_alloc_ctor(0, 2, 0); @@ -10468,7 +10085,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1() { _start: { lean_object* x_1; @@ -10476,7 +10093,7 @@ x_1 = lean_mk_string_unchecked("changes", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2() { _start: { lean_object* x_1; @@ -10484,7 +10101,7 @@ x_1 = lean_mk_string_unchecked("documentChanges", 15, 15); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3() { _start: { lean_object* x_1; @@ -10492,23 +10109,23 @@ x_1 = lean_mk_string_unchecked("changeAnnotations", 17, 17); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__1(x_3, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1; +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__3(x_6, x_5); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2; +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__3(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__5(x_9, x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__5(x_9, x_8); x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_10); @@ -10519,13 +10136,13 @@ lean_ctor_set(x_13, 1, x_12); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_4); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -10533,7 +10150,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____spec__4(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____spec__4(x_4, x_5, x_3); return x_6; } } @@ -10541,7 +10158,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_), 1, 0); return x_1; } } @@ -10553,7 +10170,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10592,7 +10209,7 @@ x_14 = lean_string_dec_eq(x_2, x_10); if (x_14 == 0) { lean_object* x_15; uint8_t x_16; -x_15 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_12, x_2, x_3); +x_15 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_12, x_2, x_3); x_16 = 0; lean_ctor_set(x_1, 3, x_15); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); @@ -10613,7 +10230,7 @@ return x_1; else { lean_object* x_18; uint8_t x_19; -x_18 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_9, x_2, x_3); +x_18 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_9, x_2, x_3); x_19 = 0; lean_ctor_set(x_1, 0, x_18); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_19); @@ -10640,7 +10257,7 @@ x_25 = lean_string_dec_eq(x_2, x_21); if (x_25 == 0) { lean_object* x_26; uint8_t x_27; lean_object* x_28; -x_26 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_23, x_2, x_3); +x_26 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_23, x_2, x_3); x_27 = 0; x_28 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_28, 0, x_20); @@ -10668,7 +10285,7 @@ return x_30; else { lean_object* x_31; uint8_t x_32; lean_object* x_33; -x_31 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_20, x_2, x_3); +x_31 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_20, x_2, x_3); x_32 = 0; x_33 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_33, 0, x_31); @@ -10699,7 +10316,7 @@ x_40 = lean_string_dec_eq(x_2, x_36); if (x_40 == 0) { lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_38, x_2, x_3); +x_41 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_38, x_2, x_3); x_42 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); if (x_42 == 0) { @@ -11390,7 +11007,7 @@ return x_1; else { lean_object* x_193; uint8_t x_194; -x_193 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_35, x_2, x_3); +x_193 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_35, x_2, x_3); x_194 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); if (x_194 == 0) { @@ -12096,7 +11713,7 @@ x_350 = lean_string_dec_eq(x_2, x_346); if (x_350 == 0) { lean_object* x_351; uint8_t x_352; -x_351 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_348, x_2, x_3); +x_351 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_348, x_2, x_3); x_352 = lean_ctor_get_uint8(x_351, sizeof(void*)*4); if (x_352 == 0) { @@ -12526,7 +12143,7 @@ return x_427; else { lean_object* x_428; uint8_t x_429; -x_428 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_345, x_2, x_3); +x_428 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_345, x_2, x_3); x_429 = lean_ctor_get_uint8(x_428, sizeof(void*)*4); if (x_429 == 0) { @@ -12943,7 +12560,7 @@ return x_502; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -12951,19 +12568,19 @@ x_4 = l_Lean_RBNode_isRed___rarg(x_1); if (x_4 == 0) { lean_object* x_5; -x_5 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_1, x_2, x_3); +x_5 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_1, x_2, x_3); return x_5; } else { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__3(x_1, x_2, x_3); +x_6 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__3(x_1, x_2, x_3); x_7 = l_Lean_RBNode_setBlack___rarg(x_6); return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -12985,7 +12602,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -13050,7 +12667,7 @@ lean_object* x_21; lean_object* x_22; x_21 = lean_ctor_get(x_17, 0); lean_inc(x_21); lean_dec(x_17); -x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -13069,10 +12686,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -13084,10 +12701,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -13098,7 +12715,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__5(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13120,7 +12737,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__5(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__5(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -13185,7 +12802,7 @@ lean_object* x_20; lean_object* x_21; x_20 = lean_ctor_get(x_16, 0); lean_inc(x_20); lean_dec(x_16); -x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -13204,10 +12821,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -13219,10 +12836,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -13233,7 +12850,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -13255,7 +12872,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -13320,7 +12937,7 @@ lean_object* x_21; lean_object* x_22; x_21 = lean_ctor_get(x_17, 0); lean_inc(x_21); lean_dec(x_17); -x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -13339,10 +12956,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -13354,10 +12971,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -13368,7 +12985,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13390,7 +13007,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -13455,7 +13072,7 @@ lean_object* x_20; lean_object* x_21; x_20 = lean_ctor_get(x_16, 0); lean_inc(x_20); lean_dec(x_16); -x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -13474,10 +13091,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -13489,10 +13106,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -13503,7 +13120,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -13525,7 +13142,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -13590,7 +13207,7 @@ lean_object* x_21; lean_object* x_22; x_21 = lean_ctor_get(x_17, 0); lean_inc(x_21); lean_dec(x_17); -x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -13609,10 +13226,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -13624,10 +13241,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -13638,7 +13255,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13660,7 +13277,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -13725,7 +13342,7 @@ lean_object* x_20; lean_object* x_21; x_20 = lean_ctor_get(x_16, 0); lean_inc(x_20); lean_dec(x_16); -x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -13744,10 +13361,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -13759,10 +13376,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -13773,7 +13390,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -13795,7 +13412,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -13860,7 +13477,7 @@ lean_object* x_21; lean_object* x_22; x_21 = lean_ctor_get(x_17, 0); lean_inc(x_21); lean_dec(x_17); -x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -13879,10 +13496,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -13894,10 +13511,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -13908,7 +13525,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13930,7 +13547,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -13995,7 +13612,7 @@ lean_object* x_20; lean_object* x_21; x_20 = lean_ctor_get(x_16, 0); lean_inc(x_20); lean_dec(x_16); -x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -14014,10 +13631,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -14029,10 +13646,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -14043,7 +13660,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -14065,7 +13682,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -14130,7 +13747,7 @@ lean_object* x_21; lean_object* x_22; x_21 = lean_ctor_get(x_17, 0); lean_inc(x_21); lean_dec(x_17); -x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_13, x_6, x_21); +x_22 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_13, x_6, x_21); x_2 = x_22; x_3 = x_8; goto _start; @@ -14149,10 +13766,10 @@ x_25 = lean_ctor_get(x_9, 0); lean_dec(x_25); x_26 = lean_unsigned_to_nat(80u); x_27 = l_Lean_Json_pretty(x_7, x_26); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); lean_ctor_set_tag(x_9, 0); lean_ctor_set(x_9, 0, x_31); @@ -14164,10 +13781,10 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean lean_dec(x_9); x_32 = lean_unsigned_to_nat(80u); x_33 = l_Lean_Json_pretty(x_7, x_32); -x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_35 = lean_string_append(x_34, x_33); lean_dec(x_33); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_37 = lean_string_append(x_35, x_36); x_38 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_38, 0, x_37); @@ -14178,7 +13795,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -14200,7 +13817,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -14265,7 +13882,7 @@ lean_object* x_20; lean_object* x_21; x_20 = lean_ctor_get(x_16, 0); lean_inc(x_20); lean_dec(x_16); -x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__2(x_12, x_5, x_20); +x_21 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__2(x_12, x_5, x_20); x_1 = x_21; x_2 = x_7; goto _start; @@ -14284,10 +13901,10 @@ x_24 = lean_ctor_get(x_8, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(80u); x_26 = l_Lean_Json_pretty(x_6, x_25); -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_30 = lean_string_append(x_28, x_29); lean_ctor_set_tag(x_8, 0); lean_ctor_set(x_8, 0, x_30); @@ -14299,10 +13916,10 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean lean_dec(x_8); x_31 = lean_unsigned_to_nat(80u); x_32 = l_Lean_Json_pretty(x_6, x_31); -x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_33 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_34 = lean_string_append(x_33, x_32); lean_dec(x_32); -x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_35 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_36 = lean_string_append(x_34, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); @@ -14313,7 +13930,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -14322,7 +13939,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -14355,7 +13972,7 @@ x_9 = lean_ctor_get(x_5, 0); lean_inc(x_9); lean_dec(x_5); x_10 = lean_box(0); -x_11 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__5(x_10, x_9); +x_11 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__5(x_10, x_9); if (lean_obj_tag(x_11) == 0) { uint8_t x_12; @@ -14441,7 +14058,7 @@ x_27 = lean_ctor_get(x_21, 0); lean_inc(x_27); lean_dec(x_21); x_28 = lean_box(0); -x_29 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(x_28, x_27); +x_29 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(x_28, x_27); if (lean_obj_tag(x_29) == 0) { uint8_t x_30; @@ -14520,7 +14137,7 @@ x_40 = lean_ctor_get(x_21, 0); lean_inc(x_40); lean_dec(x_21); x_41 = lean_box(0); -x_42 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__7(x_41, x_40); +x_42 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__7(x_41, x_40); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; @@ -14604,7 +14221,7 @@ x_56 = lean_ctor_get(x_50, 0); lean_inc(x_56); lean_dec(x_50); x_57 = lean_box(0); -x_58 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(x_57, x_56); +x_58 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(x_57, x_56); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -14683,7 +14300,7 @@ x_69 = lean_ctor_get(x_50, 0); lean_inc(x_69); lean_dec(x_50); x_70 = lean_box(0); -x_71 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__9(x_70, x_69); +x_71 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__9(x_70, x_69); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; @@ -14767,7 +14384,7 @@ x_85 = lean_ctor_get(x_79, 0); lean_inc(x_85); lean_dec(x_79); x_86 = lean_box(0); -x_87 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(x_86, x_85); +x_87 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(x_86, x_85); if (lean_obj_tag(x_87) == 0) { uint8_t x_88; @@ -14846,7 +14463,7 @@ x_98 = lean_ctor_get(x_79, 0); lean_inc(x_98); lean_dec(x_79); x_99 = lean_box(0); -x_100 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__11(x_99, x_98); +x_100 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__11(x_99, x_98); if (lean_obj_tag(x_100) == 0) { lean_object* x_101; lean_object* x_102; lean_object* x_103; @@ -14930,7 +14547,7 @@ x_114 = lean_ctor_get(x_108, 0); lean_inc(x_114); lean_dec(x_108); x_115 = lean_box(0); -x_116 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(x_115, x_114); +x_116 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(x_115, x_114); if (lean_obj_tag(x_116) == 0) { uint8_t x_117; @@ -15009,7 +14626,7 @@ x_127 = lean_ctor_get(x_108, 0); lean_inc(x_127); lean_dec(x_108); x_128 = lean_box(0); -x_129 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__13(x_128, x_127); +x_129 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__13(x_128, x_127); if (lean_obj_tag(x_129) == 0) { lean_object* x_130; lean_object* x_131; lean_object* x_132; @@ -15058,7 +14675,7 @@ return x_136; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -15123,7 +14740,7 @@ else { lean_object* x_33; lean_inc(x_6); -x_33 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952_(x_6); +x_33 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767_(x_6); if (lean_obj_tag(x_33) == 0) { lean_object* x_34; @@ -15171,7 +14788,7 @@ else lean_object* x_46; lean_dec(x_25); lean_inc(x_6); -x_46 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690_(x_6); +x_46 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505_(x_6); if (lean_obj_tag(x_46) == 0) { lean_object* x_47; @@ -15218,7 +14835,7 @@ else lean_object* x_59; lean_dec(x_25); lean_inc(x_6); -x_59 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445_(x_6); +x_59 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260_(x_6); if (lean_obj_tag(x_59) == 0) { lean_object* x_60; @@ -15274,7 +14891,7 @@ goto block_20; { lean_object* x_10; lean_dec(x_9); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660_(x_6); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475_(x_6); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; @@ -15314,7 +14931,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -15323,7 +14940,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -15331,10 +14948,10 @@ case 1: lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_5 = lean_unsigned_to_nat(80u); x_6 = l_Lean_Json_pretty(x_3, x_5); -x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_8 = lean_string_append(x_7, x_6); lean_dec(x_6); -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_10 = lean_string_append(x_8, x_9); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); @@ -15350,7 +14967,7 @@ lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; x_13 = lean_ctor_get(x_3, 0); x_14 = lean_array_size(x_13); x_15 = 0; -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(x_14, x_15, x_13); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(x_14, x_15, x_13); if (lean_obj_tag(x_16) == 0) { uint8_t x_17; @@ -15406,7 +15023,7 @@ lean_inc(x_24); lean_dec(x_3); x_25 = lean_array_size(x_24); x_26 = 0; -x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(x_25, x_26, x_24); +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(x_25, x_26, x_24); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -15463,10 +15080,10 @@ if (x_37 == 0) lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_38 = lean_ctor_get(x_3, 0); lean_dec(x_38); -x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_40 = lean_string_append(x_39, x_36); lean_dec(x_36); -x_41 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_41 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_42 = lean_string_append(x_40, x_41); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_42); @@ -15476,10 +15093,10 @@ else { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_dec(x_3); -x_43 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_43 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_44 = lean_string_append(x_43, x_36); lean_dec(x_36); -x_45 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_45 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_46 = lean_string_append(x_44, x_45); x_47 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_47, 0, x_46); @@ -15489,7 +15106,7 @@ return x_47; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -15528,7 +15145,7 @@ x_14 = lean_string_dec_eq(x_2, x_10); if (x_14 == 0) { lean_object* x_15; uint8_t x_16; -x_15 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_12, x_2, x_3); +x_15 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_12, x_2, x_3); x_16 = 0; lean_ctor_set(x_1, 3, x_15); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); @@ -15549,7 +15166,7 @@ return x_1; else { lean_object* x_18; uint8_t x_19; -x_18 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_9, x_2, x_3); +x_18 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_9, x_2, x_3); x_19 = 0; lean_ctor_set(x_1, 0, x_18); lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_19); @@ -15576,7 +15193,7 @@ x_25 = lean_string_dec_eq(x_2, x_21); if (x_25 == 0) { lean_object* x_26; uint8_t x_27; lean_object* x_28; -x_26 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_23, x_2, x_3); +x_26 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_23, x_2, x_3); x_27 = 0; x_28 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_28, 0, x_20); @@ -15604,7 +15221,7 @@ return x_30; else { lean_object* x_31; uint8_t x_32; lean_object* x_33; -x_31 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_20, x_2, x_3); +x_31 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_20, x_2, x_3); x_32 = 0; x_33 = lean_alloc_ctor(1, 4, 1); lean_ctor_set(x_33, 0, x_31); @@ -15635,7 +15252,7 @@ x_40 = lean_string_dec_eq(x_2, x_36); if (x_40 == 0) { lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_38, x_2, x_3); +x_41 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_38, x_2, x_3); x_42 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); if (x_42 == 0) { @@ -16326,7 +15943,7 @@ return x_1; else { lean_object* x_193; uint8_t x_194; -x_193 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_35, x_2, x_3); +x_193 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_35, x_2, x_3); x_194 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); if (x_194 == 0) { @@ -17032,7 +16649,7 @@ x_350 = lean_string_dec_eq(x_2, x_346); if (x_350 == 0) { lean_object* x_351; uint8_t x_352; -x_351 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_348, x_2, x_3); +x_351 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_348, x_2, x_3); x_352 = lean_ctor_get_uint8(x_351, sizeof(void*)*4); if (x_352 == 0) { @@ -17462,7 +17079,7 @@ return x_427; else { lean_object* x_428; uint8_t x_429; -x_428 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_345, x_2, x_3); +x_428 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_345, x_2, x_3); x_429 = lean_ctor_get_uint8(x_428, sizeof(void*)*4); if (x_429 == 0) { @@ -17879,7 +17496,7 @@ return x_502; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -17887,19 +17504,19 @@ x_4 = l_Lean_RBNode_isRed___rarg(x_1); if (x_4 == 0) { lean_object* x_5; -x_5 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_1, x_2, x_3); +x_5 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_1, x_2, x_3); return x_5; } else { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__18(x_1, x_2, x_3); +x_6 = l_Lean_RBNode_ins___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__18(x_1, x_2, x_3); x_7 = l_Lean_RBNode_setBlack___rarg(x_6); return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -17921,7 +17538,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -17950,7 +17567,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -17979,7 +17596,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -17988,7 +17605,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__20(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__20(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18010,7 +17627,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__20(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__20(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18039,7 +17656,7 @@ lean_object* x_12; lean_object* x_13; x_12 = lean_ctor_get(x_8, 0); lean_inc(x_12); lean_dec(x_8); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18068,7 +17685,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18077,7 +17694,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -18099,7 +17716,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -18128,7 +17745,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -18157,7 +17774,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -18166,7 +17783,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18188,7 +17805,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18217,7 +17834,7 @@ lean_object* x_12; lean_object* x_13; x_12 = lean_ctor_get(x_8, 0); lean_inc(x_12); lean_dec(x_8); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18246,7 +17863,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18255,7 +17872,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -18277,7 +17894,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -18306,7 +17923,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -18335,7 +17952,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -18344,7 +17961,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18366,7 +17983,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18395,7 +18012,7 @@ lean_object* x_12; lean_object* x_13; x_12 = lean_ctor_get(x_8, 0); lean_inc(x_12); lean_dec(x_8); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18424,7 +18041,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18433,7 +18050,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -18455,7 +18072,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -18484,7 +18101,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -18513,7 +18130,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -18522,7 +18139,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18544,7 +18161,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18573,7 +18190,7 @@ lean_object* x_12; lean_object* x_13; x_12 = lean_ctor_get(x_8, 0); lean_inc(x_12); lean_dec(x_8); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18602,7 +18219,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18611,7 +18228,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -18633,7 +18250,7 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); lean_dec(x_3); -x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27(x_1, x_2, x_5); +x_9 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27(x_1, x_2, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -18662,7 +18279,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_7); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_7); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -18691,7 +18308,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_14, 0); lean_inc(x_18); lean_dec(x_14); -x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_13, x_6, x_18); +x_19 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_13, x_6, x_18); x_2 = x_19; x_3 = x_8; goto _start; @@ -18700,7 +18317,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -18722,7 +18339,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_2, 3); lean_inc(x_7); lean_dec(x_2); -x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(x_1, x_4); +x_8 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(x_1, x_4); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; @@ -18751,7 +18368,7 @@ lean_object* x_12; lean_object* x_13; x_12 = lean_ctor_get(x_8, 0); lean_inc(x_12); lean_dec(x_8); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853_(x_6); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668_(x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -18780,7 +18397,7 @@ lean_object* x_17; lean_object* x_18; x_17 = lean_ctor_get(x_13, 0); lean_inc(x_17); lean_dec(x_13); -x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__17(x_12, x_5, x_17); +x_18 = l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__17(x_12, x_5, x_17); x_1 = x_18; x_2 = x_7; goto _start; @@ -18789,7 +18406,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -18798,7 +18415,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -18831,7 +18448,7 @@ x_9 = lean_ctor_get(x_5, 0); lean_inc(x_9); lean_dec(x_5); x_10 = lean_box(0); -x_11 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__20(x_10, x_9); +x_11 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__20(x_10, x_9); if (lean_obj_tag(x_11) == 0) { uint8_t x_12; @@ -18917,7 +18534,7 @@ x_27 = lean_ctor_get(x_21, 0); lean_inc(x_27); lean_dec(x_21); x_28 = lean_box(0); -x_29 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(x_28, x_27); +x_29 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(x_28, x_27); if (lean_obj_tag(x_29) == 0) { uint8_t x_30; @@ -18996,7 +18613,7 @@ x_40 = lean_ctor_get(x_21, 0); lean_inc(x_40); lean_dec(x_21); x_41 = lean_box(0); -x_42 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__22(x_41, x_40); +x_42 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__22(x_41, x_40); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; @@ -19080,7 +18697,7 @@ x_56 = lean_ctor_get(x_50, 0); lean_inc(x_56); lean_dec(x_50); x_57 = lean_box(0); -x_58 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(x_57, x_56); +x_58 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(x_57, x_56); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -19159,7 +18776,7 @@ x_69 = lean_ctor_get(x_50, 0); lean_inc(x_69); lean_dec(x_50); x_70 = lean_box(0); -x_71 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__24(x_70, x_69); +x_71 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__24(x_70, x_69); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; @@ -19243,7 +18860,7 @@ x_85 = lean_ctor_get(x_79, 0); lean_inc(x_85); lean_dec(x_79); x_86 = lean_box(0); -x_87 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(x_86, x_85); +x_87 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(x_86, x_85); if (lean_obj_tag(x_87) == 0) { uint8_t x_88; @@ -19322,7 +18939,7 @@ x_98 = lean_ctor_get(x_79, 0); lean_inc(x_98); lean_dec(x_79); x_99 = lean_box(0); -x_100 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__26(x_99, x_98); +x_100 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__26(x_99, x_98); if (lean_obj_tag(x_100) == 0) { lean_object* x_101; lean_object* x_102; lean_object* x_103; @@ -19406,7 +19023,7 @@ x_114 = lean_ctor_get(x_108, 0); lean_inc(x_114); lean_dec(x_108); x_115 = lean_box(0); -x_116 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(x_115, x_114); +x_116 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(x_115, x_114); if (lean_obj_tag(x_116) == 0) { uint8_t x_117; @@ -19485,7 +19102,7 @@ x_127 = lean_ctor_get(x_108, 0); lean_inc(x_127); lean_dec(x_108); x_128 = lean_box(0); -x_129 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__28(x_128, x_127); +x_129 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__28(x_128, x_127); if (lean_obj_tag(x_129) == 0) { lean_object* x_130; lean_object* x_131; lean_object* x_132; @@ -19534,7 +19151,7 @@ return x_136; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1() { _start: { lean_object* x_1; @@ -19542,39 +19159,39 @@ x_1 = lean_mk_string_unchecked("WorkspaceEdit", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5() { _start: { lean_object* x_1; @@ -19582,48 +19199,48 @@ x_1 = lean_mk_string_unchecked("changes\?", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10() { _start: { lean_object* x_1; @@ -19631,48 +19248,48 @@ x_1 = lean_mk_string_unchecked("documentChanges\?", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15() { _start: { lean_object* x_1; @@ -19680,54 +19297,54 @@ x_1 = lean_mk_string_unchecked("changeAnnotations\?", 18, 18); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -19737,7 +19354,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -19749,7 +19366,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -19763,9 +19380,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -19776,7 +19393,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -19788,7 +19405,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -19802,8 +19419,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -19814,7 +19431,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -19826,7 +19443,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -19868,61 +19485,61 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__4(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__4(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__6(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__6(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__8(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__8(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__10(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__10(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__12(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__12(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -19930,69 +19547,69 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__15(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__15(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__14(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__14(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__19(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__19(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__21(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__21(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__23(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__23(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__25(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__25(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__27(x_1, x_2, x_3); +x_4 = l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__27(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____spec__16(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____spec__16(x_1, x_2); lean_dec(x_2); return x_3; } @@ -20001,7 +19618,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_), 1, 0); return x_1; } } @@ -20451,7 +20068,7 @@ x_7 = l_Lean_Lsp_WorkspaceEdit_ofTextDocumentEdit(x_6); return x_7; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1() { _start: { lean_object* x_1; @@ -20459,19 +20076,19 @@ x_1 = lean_mk_string_unchecked("edit", 4, 4); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_3, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); lean_dec(x_1); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(x_5); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(x_5); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -20485,8 +20102,8 @@ lean_ctor_set(x_11, 1, x_9); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_4); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -20495,7 +20112,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621_), 1, 0); return x_1; } } @@ -20507,16 +20124,16 @@ x_1 = l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1() { _start: { lean_object* x_1; @@ -20524,39 +20141,39 @@ x_1 = lean_mk_string_unchecked("ApplyWorkspaceEditParams", 24, 24); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5() { _start: { lean_object* x_1; @@ -20564,95 +20181,95 @@ x_1 = lean_mk_string_unchecked("label\?", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -20662,7 +20279,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -20674,7 +20291,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -20688,8 +20305,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -20699,7 +20316,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -20711,7 +20328,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -20750,11 +20367,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -20763,7 +20380,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663_), 1, 0); return x_1; } } @@ -20775,7 +20392,7 @@ x_1 = l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1() { _start: { lean_object* x_1; @@ -20783,7 +20400,7 @@ x_1 = lean_mk_string_unchecked("languageId", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2() { _start: { lean_object* x_1; @@ -20791,7 +20408,7 @@ x_1 = lean_mk_string_unchecked("text", 4, 4); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -20806,7 +20423,7 @@ lean_inc(x_5); lean_dec(x_1); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_2); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -20816,7 +20433,7 @@ lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_11, 0, x_3); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1; x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); @@ -20826,7 +20443,7 @@ lean_ctor_set(x_14, 1, x_9); x_15 = l_Lean_JsonNumber_fromNat(x_4); x_16 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -20835,7 +20452,7 @@ lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_9); x_20 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_20, 0, x_5); -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2; +x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); @@ -20854,8 +20471,8 @@ lean_ctor_set(x_26, 1, x_25); x_27 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_27, 0, x_10); lean_ctor_set(x_27, 1, x_26); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_27, x_28); +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_27, x_28); x_30 = l_Lean_Json_mkObj(x_29); return x_30; } @@ -20864,7 +20481,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_), 1, 0); return x_1; } } @@ -20876,7 +20493,7 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentItem___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1() { _start: { lean_object* x_1; @@ -20884,188 +20501,188 @@ x_1 = lean_mk_string_unchecked("TextDocumentItem", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -21075,7 +20692,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -21087,7 +20704,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -21101,9 +20718,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -21114,7 +20731,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -21126,7 +20743,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -21140,9 +20757,9 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1; +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -21154,7 +20771,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -21166,7 +20783,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -21180,8 +20797,8 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2; -x_36 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -21193,7 +20810,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18; +x_39 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -21205,7 +20822,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -21254,7 +20871,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_), 1, 0); return x_1; } } @@ -21266,7 +20883,7 @@ x_1 = l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1() { _start: { lean_object* x_1; @@ -21274,14 +20891,14 @@ x_1 = lean_mk_string_unchecked("position", 8, 8); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -21292,8 +20909,8 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -21306,8 +20923,8 @@ lean_ctor_set(x_13, 1, x_6); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -21316,7 +20933,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentPositionParams___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089_), 1, 0); return x_1; } } @@ -21328,16 +20945,16 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1() { _start: { lean_object* x_1; @@ -21345,106 +20962,106 @@ x_1 = lean_mk_string_unchecked("TextDocumentPositionParams", 26, 26); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -21454,7 +21071,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -21466,7 +21083,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -21480,8 +21097,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -21491,7 +21108,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -21503,7 +21120,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -21542,11 +21159,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -21555,7 +21172,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentPositionParams___cl _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_), 1, 0); return x_1; } } @@ -21581,7 +21198,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToStringTextDocumentPositionParams(lean_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Lsp_instInhabitedCancelParams___closed__1; +x_3 = l_Lean_Lsp_instInhabitedLocation___closed__1; x_4 = lean_string_append(x_3, x_2); lean_dec(x_2); x_5 = l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1; @@ -21605,7 +21222,7 @@ x_15 = lean_string_append(x_14, x_3); return x_15; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1() { _start: { lean_object* x_1; @@ -21613,7 +21230,7 @@ x_1 = lean_mk_string_unchecked("language", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2() { _start: { lean_object* x_1; @@ -21621,7 +21238,7 @@ x_1 = lean_mk_string_unchecked("scheme", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3() { _start: { lean_object* x_1; @@ -21629,23 +21246,23 @@ x_1 = lean_mk_string_unchecked("pattern", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_3, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1; +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_6, x_5); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2; +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_8); x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_10); @@ -21656,8 +21273,8 @@ lean_ctor_set(x_13, 1, x_12); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_4); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -21666,7 +21283,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentFilter___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320_), 1, 0); return x_1; } } @@ -21678,7 +21295,7 @@ x_1 = l_Lean_Lsp_instToJsonDocumentFilter___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1() { _start: { lean_object* x_1; @@ -21686,39 +21303,39 @@ x_1 = lean_mk_string_unchecked("DocumentFilter", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5() { _start: { lean_object* x_1; @@ -21726,48 +21343,48 @@ x_1 = lean_mk_string_unchecked("language\?", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10() { _start: { lean_object* x_1; @@ -21775,48 +21392,48 @@ x_1 = lean_mk_string_unchecked("scheme\?", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15() { _start: { lean_object* x_1; @@ -21824,54 +21441,54 @@ x_1 = lean_mk_string_unchecked("pattern\?", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -21881,7 +21498,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -21893,7 +21510,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -21907,9 +21524,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -21920,7 +21537,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -21932,7 +21549,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -21946,8 +21563,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -21958,7 +21575,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19; +x_28 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -21970,7 +21587,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19; +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -22016,7 +21633,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDocumentFilter___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356_), 1, 0); return x_1; } } @@ -22046,7 +21663,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_6 = lean_array_uget(x_3, x_2); x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_3, x_2, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541_(x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356_(x_6); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -22092,10 +21709,10 @@ case 0: lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_2 = lean_unsigned_to_nat(80u); x_3 = l_Lean_Json_pretty(x_1, x_2); -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_5 = lean_string_append(x_4, x_3); lean_dec(x_3); -x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_7 = lean_string_append(x_5, x_6); x_8 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_8, 0, x_7); @@ -22106,10 +21723,10 @@ case 1: lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_9 = lean_unsigned_to_nat(80u); x_10 = l_Lean_Json_pretty(x_1, x_9); -x_11 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_11 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_12 = lean_string_append(x_11, x_10); lean_dec(x_10); -x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_13 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_14 = lean_string_append(x_12, x_13); x_15 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_15, 0, x_14); @@ -22138,10 +21755,10 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_23 = lean_ctor_get(x_1, 0); lean_dec(x_23); -x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_25 = lean_string_append(x_24, x_21); lean_dec(x_21); -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_27 = lean_string_append(x_25, x_26); lean_ctor_set_tag(x_1, 0); lean_ctor_set(x_1, 0, x_27); @@ -22151,10 +21768,10 @@ else { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_1); -x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_28 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_29 = lean_string_append(x_28, x_21); lean_dec(x_21); -x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_30 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_31 = lean_string_append(x_29, x_30); x_32 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_32, 0, x_31); @@ -22191,7 +21808,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -22225,18 +21842,26 @@ x_6 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSelector___spec__1( return x_6; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5727_(lean_object* x_1) { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("id", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -22245,7 +21870,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonStaticRegistrationOptions___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5727_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542_), 1, 0); return x_1; } } @@ -22257,7 +21882,7 @@ x_1 = l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1() { _start: { lean_object* x_1; @@ -22265,39 +21890,39 @@ x_1 = lean_mk_string_unchecked("StaticRegistrationOptions", 25, 25); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5() { _start: { lean_object* x_1; @@ -22305,53 +21930,53 @@ x_1 = lean_mk_string_unchecked("id\?", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -22360,7 +21985,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -22372,7 +21997,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -22405,7 +22030,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonStaticRegistrationOptions___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570_), 1, 0); return x_1; } } @@ -22417,7 +22042,7 @@ x_1 = l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -22472,7 +22097,7 @@ return x_19; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1() { _start: { lean_object* x_1; @@ -22480,18 +22105,18 @@ x_1 = lean_mk_string_unchecked("documentSelector", 16, 16); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____spec__1(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -22500,7 +22125,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651_), 1, 0); return x_1; } } @@ -22512,7 +22137,7 @@ x_1 = l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -22521,7 +22146,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -22529,10 +22154,10 @@ case 1: lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_5 = lean_unsigned_to_nat(80u); x_6 = l_Lean_Json_pretty(x_3, x_5); -x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_8 = lean_string_append(x_7, x_6); lean_dec(x_6); -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_10 = lean_string_append(x_8, x_9); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); @@ -22661,10 +22286,10 @@ if (x_37 == 0) lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; x_38 = lean_ctor_get(x_3, 0); lean_dec(x_38); -x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_39 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_40 = lean_string_append(x_39, x_36); lean_dec(x_36); -x_41 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_41 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_42 = lean_string_append(x_40, x_41); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_42); @@ -22674,10 +22299,10 @@ else { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_dec(x_3); -x_43 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1; +x_43 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1; x_44 = lean_string_append(x_43, x_36); lean_dec(x_36); -x_45 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; +x_45 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2; x_46 = lean_string_append(x_44, x_45); x_47 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_47, 0, x_46); @@ -22687,7 +22312,7 @@ return x_47; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1() { _start: { lean_object* x_1; @@ -22695,39 +22320,39 @@ x_1 = lean_mk_string_unchecked("TextDocumentRegistrationOptions", 31, 31); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5() { _start: { lean_object* x_1; @@ -22735,53 +22360,53 @@ x_1 = lean_mk_string_unchecked("documentSelector\?", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -22790,7 +22415,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -22802,7 +22427,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -22831,11 +22456,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -22844,7 +22469,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679_), 1, 0); return x_1; } } @@ -22990,7 +22615,7 @@ x_6 = lean_box(x_5); return x_6; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036_(uint8_t x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851_(uint8_t x_1) { _start: { if (x_1 == 0) @@ -23007,13 +22632,13 @@ return x_3; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851____boxed(lean_object* x_1) { _start: { uint8_t x_2; uint64_t x_3; lean_object* x_4; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851_(x_2); x_4 = lean_box_uint64(x_3); return x_4; } @@ -23022,7 +22647,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableMarkupKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851____boxed), 1, 0); return x_1; } } @@ -23188,7 +22813,7 @@ x_3 = l_Lean_Lsp_instToJsonMarkupKind(x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23200,19 +22825,19 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23224,19 +22849,19 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; @@ -23246,7 +22871,7 @@ x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -23259,34 +22884,34 @@ lean_ctor_set(x_9, 1, x_3); if (x_2 == 0) { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2; x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_13 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_11, x_12); +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_13 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_11, x_12); x_14 = l_Lean_Json_mkObj(x_13); return x_14; } else { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4; +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4; x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_9); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(x_1); lean_dec(x_1); return x_2; } @@ -23295,7 +22920,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonMarkupContent___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____boxed), 1, 0); return x_1; } } @@ -23307,7 +22932,7 @@ x_1 = l_Lean_Lsp_instToJsonMarkupContent___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -23356,7 +22981,7 @@ return x_12; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1() { _start: { lean_object* x_1; @@ -23364,39 +22989,39 @@ x_1 = lean_mk_string_unchecked("MarkupContent", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23406,64 +23031,64 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; x_2 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -23473,7 +23098,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -23485,7 +23110,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -23499,8 +23124,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -23510,7 +23135,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -23522,7 +23147,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -23565,11 +23190,11 @@ return x_30; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -23578,7 +23203,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonMarkupContent___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_), 1, 0); return x_1; } } @@ -23590,7 +23215,7 @@ x_1 = l_Lean_Lsp_instFromJsonMarkupContent___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; uint8_t x_7; @@ -23613,11 +23238,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -23628,7 +23253,7 @@ LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupContent(lean_object* x_1, le _start: { uint8_t x_3; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(x_1, x_2); return x_3; } } @@ -23643,25 +23268,25 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; uint64_t x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); x_3 = lean_ctor_get(x_1, 0); x_4 = 0; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6036_(x_2); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_5851_(x_2); x_6 = lean_uint64_mix_hash(x_4, x_5); x_7 = lean_string_hash(x_3); x_8 = lean_uint64_mix_hash(x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -23671,7 +23296,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableMarkupContent___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277____boxed), 1, 0); return x_1; } } @@ -23683,7 +23308,7 @@ x_1 = l_Lean_Lsp_instHashableMarkupContent___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1() { _start: { lean_object* x_1; @@ -23691,7 +23316,7 @@ x_1 = lean_mk_string_unchecked("token", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -23703,7 +23328,7 @@ x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_4); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1; lean_ctor_set(x_2, 1, x_6); lean_ctor_set(x_2, 0, x_7); x_8 = lean_box(0); @@ -23711,7 +23336,7 @@ x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_2); lean_ctor_set(x_9, 1, x_8); x_10 = lean_apply_1(x_1, x_5); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -23724,8 +23349,8 @@ lean_ctor_set(x_14, 1, x_8); x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_9); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_15, x_16); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_15, x_16); x_18 = l_Lean_Json_mkObj(x_17); return x_18; } @@ -23739,7 +23364,7 @@ lean_inc(x_19); lean_dec(x_2); x_21 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_21, 0, x_19); -x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1; +x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1; x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); @@ -23748,7 +23373,7 @@ x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); x_26 = lean_apply_1(x_1, x_20); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1; x_28 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_26); @@ -23761,18 +23386,18 @@ lean_ctor_set(x_30, 1, x_24); x_31 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_31, 0, x_25); lean_ctor_set(x_31, 1, x_30); -x_32 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_33 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_31, x_32); +x_32 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_33 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_31, x_32); x_34 = l_Lean_Json_mkObj(x_33); return x_34; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg), 2, 0); return x_2; } } @@ -23780,7 +23405,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams___rarg(lean_object* _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg), 2, 1); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -23793,7 +23418,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Lsp_instToJsonProgressParams___rarg), 1, return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1() { _start: { lean_object* x_1; @@ -23801,7 +23426,7 @@ x_1 = lean_mk_string_unchecked("message", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2() { _start: { lean_object* x_1; @@ -23809,7 +23434,7 @@ x_1 = lean_mk_string_unchecked("cancellable", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3() { _start: { lean_object* x_1; @@ -23817,7 +23442,7 @@ x_1 = lean_mk_string_unchecked("percentage", 10, 10); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; @@ -23839,19 +23464,19 @@ x_9 = lean_box(0); x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; -x_12 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_11, x_3); +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +x_12 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_11, x_3); x_13 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_13, 0, x_4); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2; +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_9); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_17, x_5); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3; +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_17, x_5); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_9); @@ -23864,8 +23489,8 @@ lean_ctor_set(x_21, 1, x_20); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_10); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_22, x_23); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_22, x_23); x_25 = l_Lean_Json_mkObj(x_24); return x_25; } @@ -23874,7 +23499,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456_), 1, 0); return x_1; } } @@ -23886,7 +23511,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6722_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6537_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; @@ -23910,25 +23535,25 @@ x_10 = lean_box(0); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; -x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_12, x_4); +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_12, x_4); x_14 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_14, 0, x_5); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2; +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_10); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3; -x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_18, x_6); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3; +x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_18, x_6); x_20 = lean_ctor_get(x_1, 1); lean_inc(x_20); lean_dec(x_1); x_21 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_21, 0, x_20); -x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; +x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1; x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); @@ -23950,8 +23575,8 @@ lean_ctor_set(x_28, 1, x_27); x_29 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_31 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_29, x_30); +x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_31 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_29, x_30); x_32 = l_Lean_Json_mkObj(x_31); return x_32; } @@ -23960,7 +23585,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6722_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6537_), 1, 0); return x_1; } } @@ -23972,7 +23597,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6818_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6633_(lean_object* x_1) { _start: { uint8_t x_2; @@ -23991,16 +23616,16 @@ x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_4); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_8); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -24022,16 +23647,16 @@ x_21 = lean_box(0); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_17); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1; +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_21); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_22); lean_ctor_set(x_26, 1, x_25); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } @@ -24041,7 +23666,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6818_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6633_), 1, 0); return x_1; } } @@ -24053,7 +23678,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -24100,7 +23725,7 @@ return x_12; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1() { _start: { lean_object* x_1; @@ -24108,18 +23733,18 @@ x_1 = lean_mk_string_unchecked("workDoneToken", 13, 13); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -24128,7 +23753,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689_), 1, 0); return x_1; } } @@ -24140,7 +23765,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -24149,7 +23774,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1; return x_4; } case 1: @@ -24310,7 +23935,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1() { _start: { lean_object* x_1; @@ -24318,39 +23943,39 @@ x_1 = lean_mk_string_unchecked("WorkDoneProgressParams", 22, 22); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5() { _start: { lean_object* x_1; @@ -24358,53 +23983,53 @@ x_1 = lean_mk_string_unchecked("workDoneToken\?", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24413,7 +24038,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24425,7 +24050,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24454,11 +24079,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -24467,7 +24092,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717_), 1, 0); return x_1; } } @@ -24479,7 +24104,7 @@ x_1 = l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1() { _start: { lean_object* x_1; @@ -24487,18 +24112,18 @@ x_1 = lean_mk_string_unchecked("partialResultToken", 18, 18); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -24507,7 +24132,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonPartialResultParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798_), 1, 0); return x_1; } } @@ -24519,7 +24144,7 @@ x_1 = l_Lean_Lsp_instToJsonPartialResultParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1() { _start: { lean_object* x_1; @@ -24527,39 +24152,39 @@ x_1 = lean_mk_string_unchecked("PartialResultParams", 19, 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5() { _start: { lean_object* x_1; @@ -24567,53 +24192,53 @@ x_1 = lean_mk_string_unchecked("partialResultToken\?", 19, 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24622,7 +24247,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24634,7 +24259,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24667,7 +24292,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonPartialResultParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826_), 1, 0); return x_1; } } @@ -24679,7 +24304,7 @@ x_1 = l_Lean_Lsp_instFromJsonPartialResultParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1() { _start: { lean_object* x_1; @@ -24687,13 +24312,13 @@ x_1 = lean_mk_string_unchecked("workDoneProgress", 16, 16); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905_(uint8_t x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -24704,19 +24329,19 @@ lean_ctor_set(x_6, 1, x_5); x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3; +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905_(x_2); return x_3; } } @@ -24724,7 +24349,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____boxed), 1, 0); return x_1; } } @@ -24736,7 +24361,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1() { _start: { lean_object* x_1; @@ -24744,85 +24369,85 @@ x_1 = lean_mk_string_unchecked("WorkDoneProgressOptions", 23, 23); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5; x_2 = 1; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24831,7 +24456,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24843,7 +24468,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24876,7 +24501,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943_), 1, 0); return x_1; } } @@ -24889,7 +24514,6 @@ return x_1; } } lean_object* initialize_Lean_Data_Json(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_JsonRpc(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Data_Lsp_Basic(uint8_t builtin, lean_object* w) { lean_object * res; @@ -24898,63 +24522,6 @@ _G_initialized = true; res = initialize_Lean_Data_Json(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Data_JsonRpc(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -l_Lean_Lsp_instInhabitedCancelParams___closed__1 = _init_l_Lean_Lsp_instInhabitedCancelParams___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams___closed__1); -l_Lean_Lsp_instInhabitedCancelParams___closed__2 = _init_l_Lean_Lsp_instInhabitedCancelParams___closed__2(); -lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams___closed__2); -l_Lean_Lsp_instInhabitedCancelParams = _init_l_Lean_Lsp_instInhabitedCancelParams(); -lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams); -l_Lean_Lsp_instBEqCancelParams___closed__1 = _init_l_Lean_Lsp_instBEqCancelParams___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instBEqCancelParams___closed__1); -l_Lean_Lsp_instBEqCancelParams = _init_l_Lean_Lsp_instBEqCancelParams(); -lean_mark_persistent(l_Lean_Lsp_instBEqCancelParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5); -l_Lean_Lsp_instToJsonCancelParams___closed__1 = _init_l_Lean_Lsp_instToJsonCancelParams___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCancelParams___closed__1); -l_Lean_Lsp_instToJsonCancelParams = _init_l_Lean_Lsp_instToJsonCancelParams(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCancelParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13); -l_Lean_Lsp_instFromJsonCancelParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCancelParams___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCancelParams___closed__1); -l_Lean_Lsp_instFromJsonCancelParams = _init_l_Lean_Lsp_instFromJsonCancelParams(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCancelParams); l_Lean_Lsp_instInhabitedPosition___closed__1 = _init_l_Lean_Lsp_instInhabitedPosition___closed__1(); lean_mark_persistent(l_Lean_Lsp_instInhabitedPosition___closed__1); l_Lean_Lsp_instInhabitedPosition = _init_l_Lean_Lsp_instInhabitedPosition(); @@ -24971,38 +24538,50 @@ l_Lean_Lsp_instHashablePosition___closed__1 = _init_l_Lean_Lsp_instHashablePosit lean_mark_persistent(l_Lean_Lsp_instHashablePosition___closed__1); l_Lean_Lsp_instHashablePosition = _init_l_Lean_Lsp_instHashablePosition(); lean_mark_persistent(l_Lean_Lsp_instHashablePosition); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____closed__3); l_Lean_Lsp_instToJsonPosition___closed__1 = _init_l_Lean_Lsp_instToJsonPosition___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonPosition___closed__1); l_Lean_Lsp_instToJsonPosition = _init_l_Lean_Lsp_instToJsonPosition(); lean_mark_persistent(l_Lean_Lsp_instToJsonPosition); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____closed__17); l_Lean_Lsp_instFromJsonPosition___closed__1 = _init_l_Lean_Lsp_instFromJsonPosition___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPosition___closed__1); l_Lean_Lsp_instFromJsonPosition = _init_l_Lean_Lsp_instFromJsonPosition(); @@ -25029,38 +24608,38 @@ l_Lean_Lsp_instHashableRange___closed__1 = _init_l_Lean_Lsp_instHashableRange___ lean_mark_persistent(l_Lean_Lsp_instHashableRange___closed__1); l_Lean_Lsp_instHashableRange = _init_l_Lean_Lsp_instHashableRange(); lean_mark_persistent(l_Lean_Lsp_instHashableRange); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557____closed__2); l_Lean_Lsp_instToJsonRange___closed__1 = _init_l_Lean_Lsp_instToJsonRange___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonRange___closed__1); l_Lean_Lsp_instToJsonRange = _init_l_Lean_Lsp_instToJsonRange(); lean_mark_persistent(l_Lean_Lsp_instToJsonRange); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____closed__12); l_Lean_Lsp_instFromJsonRange___closed__1 = _init_l_Lean_Lsp_instFromJsonRange___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonRange___closed__1); l_Lean_Lsp_instFromJsonRange = _init_l_Lean_Lsp_instFromJsonRange(); @@ -25075,44 +24654,46 @@ l_Lean_Lsp_instLERange = _init_l_Lean_Lsp_instLERange(); lean_mark_persistent(l_Lean_Lsp_instLERange); l_Lean_Lsp_instInhabitedLocation___closed__1 = _init_l_Lean_Lsp_instInhabitedLocation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instInhabitedLocation___closed__1); +l_Lean_Lsp_instInhabitedLocation___closed__2 = _init_l_Lean_Lsp_instInhabitedLocation___closed__2(); +lean_mark_persistent(l_Lean_Lsp_instInhabitedLocation___closed__2); l_Lean_Lsp_instInhabitedLocation = _init_l_Lean_Lsp_instInhabitedLocation(); lean_mark_persistent(l_Lean_Lsp_instInhabitedLocation); l_Lean_Lsp_instBEqLocation___closed__1 = _init_l_Lean_Lsp_instBEqLocation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instBEqLocation___closed__1); l_Lean_Lsp_instBEqLocation = _init_l_Lean_Lsp_instBEqLocation(); lean_mark_persistent(l_Lean_Lsp_instBEqLocation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895____closed__2); l_Lean_Lsp_instToJsonLocation___closed__1 = _init_l_Lean_Lsp_instToJsonLocation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonLocation___closed__1); l_Lean_Lsp_instToJsonLocation = _init_l_Lean_Lsp_instToJsonLocation(); lean_mark_persistent(l_Lean_Lsp_instToJsonLocation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____closed__12); l_Lean_Lsp_instFromJsonLocation___closed__1 = _init_l_Lean_Lsp_instFromJsonLocation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonLocation___closed__1); l_Lean_Lsp_instFromJsonLocation = _init_l_Lean_Lsp_instFromJsonLocation(); @@ -25121,194 +24702,194 @@ l_Lean_Lsp_instOrdLocation___closed__1 = _init_l_Lean_Lsp_instOrdLocation___clos lean_mark_persistent(l_Lean_Lsp_instOrdLocation___closed__1); l_Lean_Lsp_instOrdLocation = _init_l_Lean_Lsp_instOrdLocation(); lean_mark_persistent(l_Lean_Lsp_instOrdLocation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____closed__4); l_Lean_Lsp_instToJsonLocationLink___closed__1 = _init_l_Lean_Lsp_instToJsonLocationLink___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonLocationLink___closed__1); l_Lean_Lsp_instToJsonLocationLink = _init_l_Lean_Lsp_instToJsonLocationLink(); lean_mark_persistent(l_Lean_Lsp_instToJsonLocationLink); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__19); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__20); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1___closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__19); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__20); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____closed__21); l_Lean_Lsp_instFromJsonLocationLink___closed__1 = _init_l_Lean_Lsp_instFromJsonLocationLink___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonLocationLink___closed__1); l_Lean_Lsp_instFromJsonLocationLink = _init_l_Lean_Lsp_instFromJsonLocationLink(); lean_mark_persistent(l_Lean_Lsp_instFromJsonLocationLink); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____closed__3); l_Lean_Lsp_instToJsonCommand___closed__1 = _init_l_Lean_Lsp_instToJsonCommand___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonCommand___closed__1); l_Lean_Lsp_instToJsonCommand = _init_l_Lean_Lsp_instToJsonCommand(); lean_mark_persistent(l_Lean_Lsp_instToJsonCommand); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__1); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__17); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__1); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2___closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____closed__17); l_Lean_Lsp_instFromJsonCommand___closed__1 = _init_l_Lean_Lsp_instFromJsonCommand___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCommand___closed__1); l_Lean_Lsp_instFromJsonCommand = _init_l_Lean_Lsp_instFromJsonCommand(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCommand); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1655____closed__1); l_Lean_Lsp_instToJsonSnippetString___closed__1 = _init_l_Lean_Lsp_instToJsonSnippetString___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonSnippetString___closed__1); l_Lean_Lsp_instToJsonSnippetString = _init_l_Lean_Lsp_instToJsonSnippetString(); lean_mark_persistent(l_Lean_Lsp_instToJsonSnippetString); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1693____closed__8); l_Lean_Lsp_instFromJsonSnippetString___closed__1 = _init_l_Lean_Lsp_instFromJsonSnippetString___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSnippetString___closed__1); l_Lean_Lsp_instFromJsonSnippetString = _init_l_Lean_Lsp_instFromJsonSnippetString(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSnippetString); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____closed__3); l_Lean_Lsp_instToJsonTextEdit___closed__1 = _init_l_Lean_Lsp_instToJsonTextEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextEdit___closed__1); l_Lean_Lsp_instToJsonTextEdit = _init_l_Lean_Lsp_instToJsonTextEdit(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextEdit); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__19); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__19); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____closed__20); l_Lean_Lsp_instFromJsonTextEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonTextEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextEdit___closed__1); l_Lean_Lsp_instFromJsonTextEdit = _init_l_Lean_Lsp_instFromJsonTextEdit(); @@ -25323,286 +24904,286 @@ l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1 = _init_l_Lean_Lsp_instT lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1); l_Lean_Lsp_instToJsonTextDocumentIdentifier = _init_l_Lean_Lsp_instToJsonTextDocumentIdentifier(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentIdentifier); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2351____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2166____closed__6); l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1); l_Lean_Lsp_instFromJsonTextDocumentIdentifier = _init_l_Lean_Lsp_instFromJsonTextDocumentIdentifier(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentIdentifier); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____closed__1); l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1 = _init_l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1); l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier = _init_l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier(); lean_mark_persistent(l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____closed__11); l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1 = _init_l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1); l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier = _init_l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier(); lean_mark_persistent(l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2608____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2423____closed__2); l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1 = _init_l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1); l_Lean_Lsp_instToJsonTextDocumentEdit = _init_l_Lean_Lsp_instToJsonTextDocumentEdit(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentEdit); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____closed__12); l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentEdit___closed__1); l_Lean_Lsp_instFromJsonTextDocumentEdit = _init_l_Lean_Lsp_instFromJsonTextDocumentEdit(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentEdit); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2797____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2612____closed__3); l_Lean_Lsp_instToJsonChangeAnnotation___closed__1 = _init_l_Lean_Lsp_instToJsonChangeAnnotation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonChangeAnnotation___closed__1); l_Lean_Lsp_instToJsonChangeAnnotation = _init_l_Lean_Lsp_instToJsonChangeAnnotation(); lean_mark_persistent(l_Lean_Lsp_instToJsonChangeAnnotation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____closed__17); l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1 = _init_l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1); l_Lean_Lsp_instFromJsonChangeAnnotation = _init_l_Lean_Lsp_instFromJsonChangeAnnotation(); lean_mark_persistent(l_Lean_Lsp_instFromJsonChangeAnnotation); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3024____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2839____closed__2); l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1 = _init_l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1); l_Lean_Lsp_CreateFile_instToJsonOptions = _init_l_Lean_Lsp_CreateFile_instToJsonOptions(); lean_mark_persistent(l_Lean_Lsp_CreateFile_instToJsonOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3076____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_2891____closed__13); l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1 = _init_l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_CreateFile_instFromJsonOptions___closed__1); l_Lean_Lsp_CreateFile_instFromJsonOptions = _init_l_Lean_Lsp_CreateFile_instFromJsonOptions(); lean_mark_persistent(l_Lean_Lsp_CreateFile_instFromJsonOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3208____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3023____closed__2); l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1 = _init_l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1); l_Lean_Lsp_DeleteFile_instToJsonOptions = _init_l_Lean_Lsp_DeleteFile_instToJsonOptions(); lean_mark_persistent(l_Lean_Lsp_DeleteFile_instToJsonOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3075____closed__12); l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1 = _init_l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_DeleteFile_instFromJsonOptions___closed__1); l_Lean_Lsp_DeleteFile_instFromJsonOptions = _init_l_Lean_Lsp_DeleteFile_instFromJsonOptions(); lean_mark_persistent(l_Lean_Lsp_DeleteFile_instFromJsonOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3399____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3214____closed__1); l_Lean_Lsp_instToJsonCreateFile___closed__1 = _init_l_Lean_Lsp_instToJsonCreateFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonCreateFile___closed__1); l_Lean_Lsp_instToJsonCreateFile = _init_l_Lean_Lsp_instToJsonCreateFile(); lean_mark_persistent(l_Lean_Lsp_instToJsonCreateFile); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3260____closed__12); l_Lean_Lsp_instFromJsonCreateFile___closed__1 = _init_l_Lean_Lsp_instFromJsonCreateFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCreateFile___closed__1); l_Lean_Lsp_instFromJsonCreateFile = _init_l_Lean_Lsp_instFromJsonCreateFile(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCreateFile); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3630____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3445____closed__2); l_Lean_Lsp_instToJsonRenameFile___closed__1 = _init_l_Lean_Lsp_instToJsonRenameFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonRenameFile___closed__1); l_Lean_Lsp_instToJsonRenameFile = _init_l_Lean_Lsp_instToJsonRenameFile(); lean_mark_persistent(l_Lean_Lsp_instToJsonRenameFile); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3690____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3505____closed__16); l_Lean_Lsp_instFromJsonRenameFile___closed__1 = _init_l_Lean_Lsp_instFromJsonRenameFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonRenameFile___closed__1); l_Lean_Lsp_instFromJsonRenameFile = _init_l_Lean_Lsp_instFromJsonRenameFile(); @@ -25611,24 +25192,24 @@ l_Lean_Lsp_instToJsonDeleteFile___closed__1 = _init_l_Lean_Lsp_instToJsonDeleteF lean_mark_persistent(l_Lean_Lsp_instToJsonDeleteFile___closed__1); l_Lean_Lsp_instToJsonDeleteFile = _init_l_Lean_Lsp_instToJsonDeleteFile(); lean_mark_persistent(l_Lean_Lsp_instToJsonDeleteFile); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3952____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3767____closed__9); l_Lean_Lsp_instFromJsonDeleteFile___closed__1 = _init_l_Lean_Lsp_instFromJsonDeleteFile___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDeleteFile___closed__1); l_Lean_Lsp_instFromJsonDeleteFile = _init_l_Lean_Lsp_instFromJsonDeleteFile(); @@ -25647,54 +25228,54 @@ l_Lean_Lsp_instToJsonDocumentChange___closed__6 = _init_l_Lean_Lsp_instToJsonDoc lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentChange___closed__6); l_Lean_Lsp_instToJsonDocumentChange___closed__7 = _init_l_Lean_Lsp_instToJsonDocumentChange___closed__7(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentChange___closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223____closed__3); l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1 = _init_l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkspaceEdit___closed__1); l_Lean_Lsp_instToJsonWorkspaceEdit = _init_l_Lean_Lsp_instToJsonWorkspaceEdit(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkspaceEdit); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__18); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444____closed__19); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259____closed__19); l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1); l_Lean_Lsp_instFromJsonWorkspaceEdit = _init_l_Lean_Lsp_instFromJsonWorkspaceEdit(); @@ -25709,224 +25290,226 @@ l_Lean_Lsp_WorkspaceEdit_instAppend___closed__2 = _init_l_Lean_Lsp_WorkspaceEdit lean_mark_persistent(l_Lean_Lsp_WorkspaceEdit_instAppend___closed__2); l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3 = _init_l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3(); lean_mark_persistent(l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4806____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4621____closed__1); l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1 = _init_l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1); l_Lean_Lsp_instToJsonApplyWorkspaceEditParams = _init_l_Lean_Lsp_instToJsonApplyWorkspaceEditParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonApplyWorkspaceEditParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4848____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4663____closed__13); l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1 = _init_l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1); l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams = _init_l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805____closed__2); l_Lean_Lsp_instToJsonTextDocumentItem___closed__1 = _init_l_Lean_Lsp_instToJsonTextDocumentItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentItem___closed__1); l_Lean_Lsp_instToJsonTextDocumentItem = _init_l_Lean_Lsp_instToJsonTextDocumentItem(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentItem); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885____closed__18); l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1); l_Lean_Lsp_instFromJsonTextDocumentItem = _init_l_Lean_Lsp_instFromJsonTextDocumentItem(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentItem); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5274____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5089____closed__1); l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1 = _init_l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1); l_Lean_Lsp_instToJsonTextDocumentPositionParams = _init_l_Lean_Lsp_instToJsonTextDocumentPositionParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentPositionParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____closed__10); l_Lean_Lsp_instFromJsonTextDocumentPositionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentPositionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentPositionParams___closed__1); l_Lean_Lsp_instFromJsonTextDocumentPositionParams = _init_l_Lean_Lsp_instFromJsonTextDocumentPositionParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentPositionParams); l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1 = _init_l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5505____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5320____closed__3); l_Lean_Lsp_instToJsonDocumentFilter___closed__1 = _init_l_Lean_Lsp_instToJsonDocumentFilter___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentFilter___closed__1); l_Lean_Lsp_instToJsonDocumentFilter = _init_l_Lean_Lsp_instToJsonDocumentFilter(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentFilter); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__10); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__11); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__12); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__13); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__14); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__15); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__16); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__17); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__18); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5541____closed__19); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__11); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__12); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__13); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__14); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__15); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__16); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__17); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__18); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5356____closed__19); l_Lean_Lsp_instFromJsonDocumentFilter___closed__1 = _init_l_Lean_Lsp_instFromJsonDocumentFilter___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDocumentFilter___closed__1); l_Lean_Lsp_instFromJsonDocumentFilter = _init_l_Lean_Lsp_instFromJsonDocumentFilter(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDocumentFilter); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5542____closed__1); l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1 = _init_l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1); l_Lean_Lsp_instToJsonStaticRegistrationOptions = _init_l_Lean_Lsp_instToJsonStaticRegistrationOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonStaticRegistrationOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5755____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5570____closed__9); l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1); l_Lean_Lsp_instFromJsonStaticRegistrationOptions = _init_l_Lean_Lsp_instFromJsonStaticRegistrationOptions(); lean_mark_persistent(l_Lean_Lsp_instFromJsonStaticRegistrationOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5836____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5651____closed__1); l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1 = _init_l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1); l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions = _init_l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____closed__9); l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1); l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions = _init_l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions(); @@ -25953,38 +25536,38 @@ l_Lean_Lsp_instToJsonMarkupKind___closed__1 = _init_l_Lean_Lsp_instToJsonMarkupK lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupKind___closed__1); l_Lean_Lsp_instToJsonMarkupKind___closed__2 = _init_l_Lean_Lsp_instToJsonMarkupKind___closed__2(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupKind___closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974____closed__4); l_Lean_Lsp_instToJsonMarkupContent___closed__1 = _init_l_Lean_Lsp_instToJsonMarkupContent___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupContent___closed__1); l_Lean_Lsp_instToJsonMarkupContent = _init_l_Lean_Lsp_instToJsonMarkupContent(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupContent); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026____closed__10); l_Lean_Lsp_instFromJsonMarkupContent___closed__1 = _init_l_Lean_Lsp_instFromJsonMarkupContent___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonMarkupContent___closed__1); l_Lean_Lsp_instFromJsonMarkupContent = _init_l_Lean_Lsp_instFromJsonMarkupContent(); @@ -25993,14 +25576,14 @@ l_Lean_Lsp_instHashableMarkupContent___closed__1 = _init_l_Lean_Lsp_instHashable lean_mark_persistent(l_Lean_Lsp_instHashableMarkupContent___closed__1); l_Lean_Lsp_instHashableMarkupContent = _init_l_Lean_Lsp_instHashableMarkupContent(); lean_mark_persistent(l_Lean_Lsp_instHashableMarkupContent); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6531____rarg___closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6641____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6346____rarg___closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6456____closed__3); l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressReport = _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport(); @@ -26013,84 +25596,84 @@ l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1 = _init_l_Lean_Lsp_instToJs lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressEnd = _init_l_Lean_Lsp_instToJsonWorkDoneProgressEnd(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressEnd); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressParams = _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____closed__9); l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1); l_Lean_Lsp_instFromJsonWorkDoneProgressParams = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6983____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6798____closed__1); l_Lean_Lsp_instToJsonPartialResultParams___closed__1 = _init_l_Lean_Lsp_instToJsonPartialResultParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonPartialResultParams___closed__1); l_Lean_Lsp_instToJsonPartialResultParams = _init_l_Lean_Lsp_instToJsonPartialResultParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonPartialResultParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7011____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6826____closed__9); l_Lean_Lsp_instFromJsonPartialResultParams___closed__1 = _init_l_Lean_Lsp_instFromJsonPartialResultParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPartialResultParams___closed__1); l_Lean_Lsp_instFromJsonPartialResultParams = _init_l_Lean_Lsp_instFromJsonPartialResultParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPartialResultParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7090____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6905____closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressOptions = _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7128____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6943____closed__8); l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1); l_Lean_Lsp_instFromJsonWorkDoneProgressOptions = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions(); diff --git a/stage0/stdlib/Lean/Data/Lsp/CancelParams.c b/stage0/stdlib/Lean/Data/Lsp/CancelParams.c new file mode 100644 index 000000000000..27bade501d73 --- /dev/null +++ b/stage0/stdlib/Lean/Data/Lsp/CancelParams.c @@ -0,0 +1,603 @@ +// Lean compiler output +// Module: Lean.Data.Lsp.CancelParams +// Imports: Lean.Data.JsonRpc +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCancelParams; +static lean_object* l_Lean_Lsp_instToJsonCancelParams___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7; +LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCancelParams; +lean_object* l_Lean_Json_mkObj(lean_object*); +LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6; +static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28____boxed(lean_object*, lean_object*); +lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCancelParams; +lean_object* lean_array_to_list(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCancelParams; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2; +static lean_object* l_Lean_Lsp_instBEqCancelParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3; +lean_object* l_List_foldl___at_Array_appendList___spec__1___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__2; +lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4; +uint8_t l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_36_(lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3; +static lean_object* l_Lean_Lsp_instFromJsonCancelParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10; +static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Lsp_instInhabitedCancelParams___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Lsp_instInhabitedCancelParams() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instInhabitedCancelParams___closed__2; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28_(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_36_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Lsp_instBEqCancelParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_28____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instBEqCancelParams() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instBEqCancelParams___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = lean_array_to_list(x_2); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +lean_dec(x_1); +x_6 = l_List_foldl___at_Array_appendList___spec__1___rarg(x_2, x_4); +x_1 = x_5; +x_2 = x_6; +goto _start; +} +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("id", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_ctor_set_tag(x_1, 3); +x_4 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_1); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_2); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_2); +x_8 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_7, x_8); +x_10 = l_Lean_Json_mkObj(x_9); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_dec(x_1); +x_12 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_12, 0, x_11); +x_13 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_2); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_2); +x_17 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_16, x_17); +x_19 = l_Lean_Json_mkObj(x_18); +return x_19; +} +} +case 1: +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_1); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_ctor_set_tag(x_1, 2); +x_21 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_1); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_2); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_2); +x_25 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_24, x_25); +x_27 = l_Lean_Json_mkObj(x_26); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_28 = lean_ctor_get(x_1, 0); +lean_inc(x_28); +lean_dec(x_1); +x_29 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_2); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_2); +x_34 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_33, x_34); +x_36 = l_Lean_Json_mkObj(x_35); +return x_36; +} +} +default: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5; +x_38 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1; +x_39 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_37, x_38); +x_40 = l_Lean_Json_mkObj(x_39); +return x_40; +} +} +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCancelParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCancelParams() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instToJsonCancelParams___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lsp", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("CancelParams", 12, 12); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2; +x_3 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4; +x_2 = 1; +x_3 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9; +x_2 = 1; +x_3 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2; +x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +x_6 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13; +x_7 = lean_string_append(x_6, x_5); +lean_dec(x_5); +lean_ctor_set(x_3, 0, x_7); +return x_3; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13; +x_10 = lean_string_append(x_9, x_8); +lean_dec(x_8); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +return x_3; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_3, 0); +lean_inc(x_13); +lean_dec(x_3); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonCancelParams___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonCancelParams() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instFromJsonCancelParams___closed__1; +return x_1; +} +} +lean_object* initialize_Lean_Data_JsonRpc(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Data_Lsp_CancelParams(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Data_JsonRpc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Lsp_instInhabitedCancelParams___closed__1 = _init_l_Lean_Lsp_instInhabitedCancelParams___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams___closed__1); +l_Lean_Lsp_instInhabitedCancelParams___closed__2 = _init_l_Lean_Lsp_instInhabitedCancelParams___closed__2(); +lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams___closed__2); +l_Lean_Lsp_instInhabitedCancelParams = _init_l_Lean_Lsp_instInhabitedCancelParams(); +lean_mark_persistent(l_Lean_Lsp_instInhabitedCancelParams); +l_Lean_Lsp_instBEqCancelParams___closed__1 = _init_l_Lean_Lsp_instBEqCancelParams___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instBEqCancelParams___closed__1); +l_Lean_Lsp_instBEqCancelParams = _init_l_Lean_Lsp_instBEqCancelParams(); +lean_mark_persistent(l_Lean_Lsp_instBEqCancelParams); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__1); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__2); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__3); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__4); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____closed__5); +l_Lean_Lsp_instToJsonCancelParams___closed__1 = _init_l_Lean_Lsp_instToJsonCancelParams___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instToJsonCancelParams___closed__1); +l_Lean_Lsp_instToJsonCancelParams = _init_l_Lean_Lsp_instToJsonCancelParams(); +lean_mark_persistent(l_Lean_Lsp_instToJsonCancelParams); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__1); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__2); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__3); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__4); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__5); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__6); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__7); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__8); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__9); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__10); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__11); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__12); +l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13 = _init_l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124____closed__13); +l_Lean_Lsp_instFromJsonCancelParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCancelParams___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonCancelParams___closed__1); +l_Lean_Lsp_instFromJsonCancelParams = _init_l_Lean_Lsp_instFromJsonCancelParams(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonCancelParams); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Data/Lsp/Client.c b/stage0/stdlib/Lean/Data/Lsp/Client.c index 51a4f39710a6..d665aa401a6a 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Client.c +++ b/stage0/stdlib/Lean/Data/Lsp/Client.c @@ -27,7 +27,6 @@ static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistr LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_296____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonRegistrationParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____closed__18; -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____closed__19; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_258____spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -44,7 +43,6 @@ static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistr static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_296____spec__1___closed__1; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____lambda__1(lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__5; static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_296____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRegistration; @@ -68,6 +66,7 @@ static lean_object* l_Lean_Lsp_instToJsonRegistrationParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRegistration; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_296_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____spec__1(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); @@ -89,6 +88,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Cli static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____boxed(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_100____closed__6; static lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__4; static lean_object* _init_l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__1() { @@ -199,7 +199,7 @@ x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_9); lean_ctor_set(x_16, 1, x_15); x_17 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__3; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } @@ -225,7 +225,7 @@ x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_9); lean_ctor_set(x_26, 1, x_25); x_27 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__3; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } @@ -505,7 +505,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_fromJsonRe lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -543,7 +543,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -702,7 +702,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); x_11 = l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistration____x40_Lean_Data_Lsp_Client___hyg_34____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } diff --git a/stage0/stdlib/Lean/Data/Lsp/CodeActions.c b/stage0/stdlib/Lean/Data/Lsp/CodeActions.c index e020972f5b90..209fa40073ed 100644 --- a/stage0/stdlib/Lean/Data/Lsp/CodeActions.c +++ b/stage0/stdlib/Lean/Data/Lsp/CodeActions.c @@ -36,6 +36,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCodeActionLiteralSupportValueSet; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__45; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__41; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2351_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCodeActionTriggerKind(uint8_t); @@ -63,8 +64,10 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__5; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__31; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_839____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__24; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__3; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__1; @@ -74,12 +77,12 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__25; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_839_(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____closed__5; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__22; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__29; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__26; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionLiteralSupportValueSet____x40_Lean_Data_Lsp_CodeActions___hyg_1754____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__17; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__24; @@ -98,7 +101,6 @@ static lean_object* l_Lean_Lsp_instFromJsonCodeActionDisabled___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_293____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__2; lean_object* l_Lean_Json_getBool_x3f(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__9; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__10; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__1; @@ -134,25 +136,24 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__9; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__19; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____spec__1___boxed(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____spec__4___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonCodeActionTriggerKind___closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonResolveSupport____x40_Lean_Data_Lsp_CodeActions___hyg_1636____closed__3; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__27; LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion___rarg___lambda__1(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__12; -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__30; static lean_object* l_Lean_Lsp_instFromJsonCodeActionOptions___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonResolveSupport____x40_Lean_Data_Lsp_CodeActions___hyg_1703____spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__8; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__7; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__5; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_293____closed__1; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__3; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____spec__2(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonResolveSupport___closed__1; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____closed__4; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); @@ -189,7 +190,7 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__6; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionLiteralSupportValueSet____x40_Lean_Data_Lsp_CodeActions___hyg_1754____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion___rarg___lambda__1___boxed(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__12; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__4(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__9; @@ -209,13 +210,13 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l_Lean_Lsp_instToJsonCodeActionTriggerKind___closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__39; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__16; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__21; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__15; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__2; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____spec__5(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__41; static lean_object* l_Lean_Lsp_instToJsonCodeActionTriggerKind___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionLiteralSupportValueSet____x40_Lean_Data_Lsp_CodeActions___hyg_1754____closed__9; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__14; @@ -235,14 +236,16 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__32; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCodeActionLiteralSupportValueSet; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____closed__16; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__42; static lean_object* l_Lean_Lsp_instFromJsonCodeActionTriggerKind___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_293____spec__2___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__13; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_839____spec__1___boxed(lean_object*, lean_object*); @@ -262,7 +265,6 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_699____closed__9; static lean_object* l_Lean_Lsp_instToJsonCodeActionTriggerKind___closed__3; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____spec__2___boxed(lean_object*, lean_object*); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____spec__3(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__7; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__33; @@ -296,7 +298,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_CodeActionTriggerKind_noConfusion(lean_objec static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__1; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_885____closed__7; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__27; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__6; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____closed__27; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__9; @@ -344,7 +345,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1205____closed__30; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2050____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_148____spec__2___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonResolveSupport____x40_Lean_Data_Lsp_CodeActions___hyg_1636_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_839____spec__1(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -2100,7 +2100,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2138,7 +2138,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__11; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2177,7 +2177,7 @@ lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__17; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -2217,7 +2217,7 @@ lean_inc(x_34); lean_dec(x_25); x_35 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__22; lean_inc(x_1); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -2362,14 +2362,14 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__11; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_6, x_5); +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_8); x_10 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__17; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -2380,7 +2380,7 @@ lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); x_14 = lean_ctor_get(x_1, 3); lean_inc(x_14); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_14); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_14); x_16 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__22; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); @@ -3422,7 +3422,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -3450,7 +3450,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -3525,11 +3525,11 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__11; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6874____spec__1(x_6, x_5); +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6689____spec__1(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); x_9 = lean_alloc_ctor(3, 1, 0); @@ -3571,7 +3571,7 @@ x_32 = lean_ctor_get(x_1, 9); lean_inc(x_32); lean_dec(x_1); x_33 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1131____closed__7; -x_34 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_33, x_32); +x_34 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_33, x_32); lean_dec(x_32); x_35 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_35, 0, x_34); @@ -4155,7 +4155,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -4206,7 +4206,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4444_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4259_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -4325,7 +4325,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -4376,7 +4376,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -4950,7 +4950,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4988,7 +4988,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_390____closed__11; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6902____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6717____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c b/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c index 5c5aa0e4d93a..49e76388f362 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c +++ b/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c @@ -27,20 +27,19 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__12; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__3(lean_object*, lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__16; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__19; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticWith(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___boxed(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag___closed__4; lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg___closed__1; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553_(lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__4; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__21; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag(lean_object*); @@ -56,7 +55,6 @@ LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_Diagnostics_0_ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__40; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instOrdDiagnosticCode; -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__28; static lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____spec__2(lean_object*); @@ -69,14 +67,14 @@ static lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag___closed__2; static lean_object* l_Lean_Lsp_instInhabitedDiagnosticRelatedInformation___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____rarg(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag___closed__3; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__1___closed__5; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__22; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__43; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__5; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__6; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____closed__7; static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__5; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__3___closed__1; @@ -93,6 +91,7 @@ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDi static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__35; LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticTag_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__1(lean_object*, lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Data_Json_Basic_0__Lean_Json_beq_x27(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____closed__14; @@ -110,7 +109,6 @@ static lean_object* l_Lean_Lsp_instBEqPublishDiagnosticsParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqDiagnosticCode; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticTag____x40_Lean_Data_Lsp_Diagnostics___hyg_628_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455_(lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticWith_fullRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__7___boxed(lean_object*, lean_object*); @@ -127,15 +125,16 @@ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDi LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2324____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity(uint8_t); LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____spec__2___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPublishDiagnosticsParams; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticTag____x40_Lean_Data_Lsp_Diagnostics___hyg_646_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedDiagnosticWith(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instBEqDiagnosticCode___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg(uint8_t, uint8_t, lean_object*); lean_object* lean_nat_to_int(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity___closed__2; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqDiagnosticSeverity; LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedDiagnosticTag; @@ -186,7 +185,6 @@ static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Dia LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_DiagnosticWith_ordUserVisible____x40_Lean_Data_Lsp_Diagnostics___hyg_2056_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticWith___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqPublishDiagnosticsParams; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__34; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_1062_(lean_object*, lean_object*); @@ -202,13 +200,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDi static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__31; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__8; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____closed__15; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__39; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__17; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__6___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__4; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticSeverity____x40_Lean_Data_Lsp_Diagnostics___hyg_39____boxed(lean_object*, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__16; static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedDiagnosticRelatedInformation; @@ -252,13 +253,12 @@ static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1254____rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__10; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____closed__2; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__38; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticTag(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticCode(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_1062____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticRelatedInformation; -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_DiagnosticWith_UserVisible_ofDiagnostic___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedDiagnosticSeverity; @@ -271,7 +271,6 @@ lean_object* l_Lean_Json_getNat_x3f(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470____spec__4(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticCode____x40_Lean_Data_Lsp_Diagnostics___hyg_327____boxed(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____spec__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity(lean_object*); static lean_object* l_Lean_Lsp_instInhabitedPublishDiagnosticsParams___closed__2; @@ -296,11 +295,10 @@ lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__26; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticSeverity____x40_Lean_Data_Lsp_Diagnostics___hyg_39_(uint8_t, uint8_t); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(lean_object*, lean_object*); uint8_t lean_int_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__23; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__17; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_904____closed__2; lean_object* lean_array_get_size(lean_object*); @@ -319,6 +317,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_Diagn static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__25; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_904_(lean_object*); lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__8; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1553____rarg___closed__30; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_830_(lean_object*, lean_object*); @@ -335,6 +334,8 @@ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPu static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____closed__3; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2324____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticSeverity____x40_Lean_Data_Lsp_Diagnostics___hyg_21_(uint8_t, uint8_t); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_956____closed__3; @@ -1804,7 +1805,7 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(x_3, x_5); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_821_(x_3, x_5); if (x_7 == 0) { uint8_t x_8; @@ -1877,7 +1878,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJso lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(x_2); x_4 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_904____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -1931,7 +1932,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(x_3); return x_4; } } @@ -2252,7 +2253,7 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(x_3, x_5); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1053_(x_3, x_5); x_8 = lean_string_dec_lt(x_4, x_6); if (x_8 == 0) { @@ -2394,7 +2395,7 @@ else lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_2, 0); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_6, x_7); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_6, x_7); return x_8; } } @@ -2751,7 +2752,7 @@ lean_inc(x_20); x_21 = lean_ctor_get(x_3, 8); lean_inc(x_21); lean_dec(x_3); -x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_13); +x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_13); lean_dec(x_13); lean_dec(x_4); if (x_22 == 0) @@ -3400,6 +3401,32 @@ return x_19; } } } +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +} +} static lean_object* _init_l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1() { _start: { @@ -3470,7 +3497,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJso lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_5 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); @@ -3482,7 +3509,7 @@ lean_ctor_set(x_8, 1, x_7); x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); x_10 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__2; -x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_10, x_9); +x_11 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(x_10, x_9); x_12 = lean_ctor_get(x_2, 2); lean_inc(x_12); x_13 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__3; @@ -3495,7 +3522,7 @@ x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJs x_18 = lean_ctor_get(x_2, 4); lean_inc(x_18); x_19 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__5; -x_20 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_19, x_18); +x_20 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_19, x_18); x_21 = lean_ctor_get(x_2, 5); lean_inc(x_21); x_22 = lean_apply_1(x_1, x_21); @@ -3518,7 +3545,7 @@ x_32 = lean_ctor_get(x_2, 8); lean_inc(x_32); lean_dec(x_2); x_33 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__8; -x_34 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_33, x_32); +x_34 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_33, x_32); lean_dec(x_32); x_35 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_35, 0, x_34); @@ -3594,6 +3621,15 @@ x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Ls return x_6; } } +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticWith___rarg(lean_object* x_1) { _start: { @@ -5048,7 +5084,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJ lean_object* x_3; lean_object* x_4; x_3 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1; lean_inc(x_2); -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_2, x_3); +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_2, x_3); if (lean_obj_tag(x_4) == 0) { uint8_t x_5; @@ -5087,7 +5123,7 @@ lean_inc(x_13); lean_dec(x_4); x_14 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__2; lean_inc(x_2); -x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_2, x_14); +x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_2, x_14); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; @@ -5210,7 +5246,7 @@ lean_inc(x_46); lean_dec(x_37); x_47 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__5; lean_inc(x_2); -x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_2, x_47); +x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_2, x_47); if (lean_obj_tag(x_48) == 0) { uint8_t x_49; @@ -5603,7 +5639,7 @@ lean_inc(x_18); x_19 = lean_ctor_get(x_3, 7); lean_inc(x_19); lean_dec(x_3); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_4, x_12); +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_4, x_12); lean_dec(x_12); lean_dec(x_4); x_21 = lean_apply_2(x_1, x_9, x_17); @@ -5646,7 +5682,7 @@ lean_dec(x_5); x_125 = lean_ctor_get(x_13, 0); lean_inc(x_125); lean_dec(x_13); -x_126 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_124, x_125); +x_126 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_124, x_125); lean_dec(x_125); lean_dec(x_124); x_23 = x_126; @@ -6422,7 +6458,7 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_2, 8); lean_inc(x_20); lean_dec(x_2); -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_3, x_12); +x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_3, x_12); lean_dec(x_12); lean_dec(x_3); if (x_21 == 0) @@ -6804,7 +6840,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJso lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_2); x_4 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -6816,7 +6852,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); x_9 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_9, x_8); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(x_9, x_8); x_11 = lean_ctor_get(x_1, 2); lean_inc(x_11); x_12 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__3; @@ -6829,7 +6865,7 @@ x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJs x_17 = lean_ctor_get(x_1, 4); lean_inc(x_17); x_18 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__5; -x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_18, x_17); +x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_18, x_17); x_20 = lean_ctor_get(x_1, 5); lean_inc(x_20); x_21 = lean_alloc_ctor(3, 1, 0); @@ -6853,7 +6889,7 @@ x_31 = lean_ctor_get(x_1, 8); lean_inc(x_31); lean_dec(x_1); x_32 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__8; -x_33 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_32, x_31); +x_33 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_32, x_31); lean_dec(x_31); x_34 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_34, 0, x_33); @@ -7191,7 +7227,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7229,7 +7265,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7349,7 +7385,7 @@ lean_inc(x_45); lean_dec(x_36); x_46 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____rarg___closed__5; lean_inc(x_1); -x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_46); +x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_46); if (lean_obj_tag(x_47) == 0) { uint8_t x_48; @@ -7886,7 +7922,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2414____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; diff --git a/stage0/stdlib/Lean/Data/Lsp/Extra.c b/stage0/stdlib/Lean/Data/Lsp/Extra.c index 0b4afba07b51..938bdfbddaf7 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Extra.c +++ b/stage0/stdlib/Lean/Data/Lsp/Extra.c @@ -37,6 +37,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRan static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__12; lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__8; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2944_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonPlainGoalParams___closed__1; @@ -76,11 +77,11 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallP static lean_object* l_Lean_Lsp_instFromJsonLeanDidOpenTextDocumentParams___closed__1; uint8_t l_Lean_Name_isAnonymous(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____spec__2(size_t, size_t, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__17; static lean_object* l_Lean_Lsp_instToJsonLeanFileProgressKind___closed__4; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2838____closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__4; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1756____closed__3; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____spec__1___closed__1; static lean_object* l_Lean_Lsp_instFromJsonWaitForDiagnostics___closed__1; @@ -99,14 +100,12 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFile static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1873_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2838____closed__7; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1586____closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__6; static lean_object* l_Lean_Lsp_instToJsonPlainTermGoalParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__4; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__3; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__20; lean_object* l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__6; @@ -119,12 +118,12 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFile static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_995____spec__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedDependencyBuildMode; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__19; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____closed__6; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonPlainTermGoalParams; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRpcCallParams; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__14; @@ -148,6 +147,7 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallP LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_133_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1756____closed__2; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__1___closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; static lean_object* l_Lean_Lsp_instToJsonLeanDidOpenTextDocumentParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__15; @@ -167,6 +167,7 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLineRang lean_object* lean_nat_to_int(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2360____closed__1; static lean_object* l_Lean_Lsp_instFromJsonLeanFileProgressKind___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_133____closed__3; static lean_object* l_Lean_Lsp_instFromJsonRpcConnectParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889____closed__8; @@ -177,7 +178,6 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallP static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1586____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_995_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__7; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889____spec__1___closed__1; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____spec__1___boxed(lean_object*, lean_object*); @@ -199,10 +199,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_LeanFileProgressKind_toCtorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____closed__11; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_LeanFileProgressKind_noConfusion(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__4; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__5; @@ -226,6 +224,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLineRange; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2360____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRpcConnected; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1873____spec__1___closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__8; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__10; @@ -236,7 +235,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_DependencyBuildMode_toCtorIdx(uint8_t); static lean_object* l_Lean_Lsp_instToJsonLeanFileProgressKind___closed__2; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_816____closed__3; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWaitForDiagnosticsParams; @@ -261,6 +259,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_LeanFileProgressKind_toCtorIdx(uint8_t); static lean_object* l_Lean_Lsp_instInhabitedLineRange___closed__1; static lean_object* l_Lean_Lsp_instToJsonLeanFileProgressKind___closed__1; static lean_object* l_Lean_Lsp_instToJsonPlainTermGoal___closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWaitForDiagnostics(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__1; static lean_object* l_Lean_Lsp_instFromJsonDependencyBuildMode___closed__1; @@ -269,11 +268,13 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPlainTermGoalParams; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcKeepAliveParams____x40_Lean_Data_Lsp_Extra___hyg_2591____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____lambda__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1586_(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_133____closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__8; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__3; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710_(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRpcConnected; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__10; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____closed__16; @@ -306,7 +307,6 @@ static lean_object* l_Lean_Lsp_instFromJsonPlainGoalParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPlainGoalParams; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889____closed__4; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____lambda__3___closed__2; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__16; lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1940_(uint64_t); @@ -373,9 +373,9 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2237_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1873____spec__1___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instBEqLeanFileProgressKind___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__11; lean_object* l_Lean_bignumToJson(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPlainTermGoal; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1873____closed__7; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889____spec__1___boxed(lean_object*, lean_object*); @@ -387,6 +387,7 @@ lean_object* lean_array_mk(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1408____closed__2; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__10; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2360____closed__10; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_995____spec__1(size_t, size_t, lean_object*); @@ -403,7 +404,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla static lean_object* l_Lean_Lsp_instToJsonWaitForDiagnostics___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1756____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWaitForDiagnosticsParams; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonDependencyBuildMode____x40_Lean_Data_Lsp_Extra___hyg_11____closed__1; lean_object* lean_string_append(lean_object*, lean_object*); @@ -426,12 +426,12 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForD LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLeanFileProgressParams; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonRpcKeepAliveParams____x40_Lean_Data_Lsp_Extra___hyg_2697_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__6; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcKeepAliveParams____x40_Lean_Data_Lsp_Extra___hyg_2591____closed__5; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2360____closed__13; static lean_object* l_Lean_Lsp_instToJsonLineRange___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_816_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1344____spec__1(size_t, size_t, lean_object*); @@ -1475,7 +1475,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanD lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -1661,7 +1661,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWai lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1698,7 +1698,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2604,7 +2604,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLea lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2809,7 +2809,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanF lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -3123,7 +3123,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLea lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3289,7 +3289,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanF lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -3468,7 +3468,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3505,7 +3505,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -3588,7 +3588,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlain lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -3600,7 +3600,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -3918,7 +3918,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1238____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4278,7 +4278,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4315,7 +4315,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -4398,7 +4398,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlain lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -4410,7 +4410,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -4561,7 +4561,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPla lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1586____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4598,7 +4598,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -4694,7 +4694,7 @@ x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_10 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -4731,7 +4731,7 @@ x_23 = lean_box(0); x_24 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_19); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_19); x_26 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_710____closed__1; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); @@ -4833,7 +4833,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpc { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5626,7 +5626,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpc lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5664,7 +5664,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -5879,7 +5879,7 @@ x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_3); x_5 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_203____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); @@ -5891,7 +5891,7 @@ lean_ctor_set(x_8, 1, x_7); x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_dec(x_2); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_9); x_11 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1059____closed__7; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -6243,7 +6243,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpc lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -6622,7 +6622,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpc lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_371____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7214,7 +7214,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLin lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7251,7 +7251,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2780____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Data/Lsp/InitShutdown.c b/stage0/stdlib/Lean/Data/Lsp/InitShutdown.c index f9023328c474..89e4465c0a6c 100644 --- a/stage0/stdlib/Lean/Data/Lsp/InitShutdown.c +++ b/stage0/stdlib/Lean/Data/Lsp/InitShutdown.c @@ -44,7 +44,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_InitializedParams_toCtorIdx___boxed(lean_obj static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__10; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_70____closed__9; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__6(lean_object*, lean_object*); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__6; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__10; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_70____closed__10; @@ -56,6 +55,7 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonC LEAN_EXPORT lean_object* l_Lean_Lsp_Trace_hasToJson(uint8_t); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializationOptions____x40_Lean_Data_Lsp_InitShutdown___hyg_343____closed__3; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__12; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonInitializeResult; static lean_object* l_Lean_Lsp_instToJsonClientInfo___closed__1; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__3; @@ -76,7 +76,6 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonI lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_InitializedParams_toCtorIdx(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_Trace_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonServerInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_800_(lean_object*); @@ -119,12 +118,14 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonIni static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_70____closed__1; static lean_object* l_Lean_Lsp_instFromJsonServerInfo___closed__1; static lean_object* l_Lean_Lsp_Trace_hasToJson___closed__1; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_InitializeParams_editDelay(lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__5; lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1171_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__9; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__9; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_928____closed__1; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__3; @@ -149,7 +150,6 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonC LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__2; static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_InitializedParams_noConfusion___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970____spec__2(lean_object*, lean_object*); @@ -165,8 +165,8 @@ static lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonC LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonServerInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_758_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_InitializedParams_noConfusion___rarg(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__6___closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_fromJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_970_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_Trace_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_928_(lean_object*); @@ -253,7 +253,7 @@ x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); x_9 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_4); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); @@ -284,7 +284,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); x_23 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_17); +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_21); @@ -535,7 +535,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -934,7 +934,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializationOptions____x40_Lean_Data_Lsp_InitShutdown___hyg_311____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440____spec__1(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); lean_dec(x_1); @@ -1114,7 +1114,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_from lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializationOptions____x40_Lean_Data_Lsp_InitShutdown___hyg_311____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1507,7 +1507,7 @@ x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJs x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); x_9 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__3; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_8); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_8); x_11 = lean_ctor_get(x_1, 3); lean_inc(x_11); x_12 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeParams____x40_Lean_Data_Lsp_InitShutdown___hyg_523____closed__4; @@ -4191,7 +4191,7 @@ x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); x_9 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_4); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); @@ -4222,7 +4222,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); x_23 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_17); +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_21); @@ -4375,7 +4375,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonClientInfo____x40_Lean_Data_Lsp_InitShutdown___hyg_28____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Data/Lsp/Internal.c b/stage0/stdlib/Lean/Data/Lsp/Internal.c index aa45fbde2173..32d55b6945f8 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Internal.c +++ b/stage0/stdlib/Lean/Data/Lsp/Internal.c @@ -30,6 +30,7 @@ static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanI static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__4; lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__1___closed__2; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonModuleRefs___spec__1(size_t, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_RefInfo_instInhabitedLocation___closed__2; @@ -58,12 +59,12 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLeanIleanInfoParams; static lean_object* l_Lean_Lsp_RefIdent_instFromJsonRefIdentJsonRepr___closed__1; uint64_t lean_string_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonModuleRefs(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonLeanImportClosureParams___closed__1; extern lean_object* l_Lean_instInhabitedJson; lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__9; uint8_t lean_string_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLeanStaleDependencyParams; static lean_object* l_Lean_Lsp_instInhabitedRefIdent___closed__1; static lean_object* l_Lean_Lsp_RefIdent_instToJsonRefIdentJsonRepr___closed__1; @@ -75,7 +76,6 @@ static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanI lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__2; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2___closed__2; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__7; @@ -104,14 +104,12 @@ lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2161____spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRefInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instInhabitedRefIdent___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonModuleRefs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_foldM___at_Lean_Lsp_instFromJsonModuleRefs___spec__6___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__1___boxed(lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__3; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_2029____spec__5(lean_object*, lean_object*, lean_object*); @@ -125,6 +123,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_fromJson_x3f(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__16; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____closed__5; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_2029____spec__2(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_beqRefIdent____x40_Lean_Data_Lsp_Internal___hyg_45_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__9; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2___closed__1; @@ -153,6 +152,7 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Lsp_instToJsonModuleRefs___ lean_object* l_Lean_Json_parseTagged(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2(size_t, size_t, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__2; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__2___closed__1; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__1; @@ -196,6 +196,7 @@ static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanI static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__8; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Lsp_instFromJsonModuleRefs___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__19; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__3___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonRefInfo___closed__1; @@ -203,13 +204,12 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_instFromJson; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_RefInfo_instInhabitedLocation___closed__1; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____lambda__2___closed__3; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__2; static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefIdent_fromJsonRefIdentJsonRepr____x40_Lean_Data_Lsp_Internal___hyg_270____closed__3; -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2094____closed__9; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l_Lean_Lsp_RefInfo_instInhabitedLocation___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2094_(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -226,13 +226,13 @@ lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Lsp_instFromJsonModuleRefs___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2161_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__10; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_instToJson; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2094____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanImportClosureParams____x40_Lean_Data_Lsp_Internal___hyg_2094____spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonLeanIleanInfoParams___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__3(size_t, size_t, size_t, lean_object*); size_t lean_usize_land(size_t, size_t); @@ -1780,7 +1780,7 @@ x_8 = lean_box(0); x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_11 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -1788,7 +1788,7 @@ lean_ctor_set(x_12, 1, x_10); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_8); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_15 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__3; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -1806,7 +1806,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_9); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -2873,7 +2873,7 @@ x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = lean_array_mk(x_6); x_10 = lean_array_size(x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_1, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_1, x_9); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = 1; @@ -2996,7 +2996,7 @@ lean_object* x_32; lean_object* x_33; size_t x_34; lean_object* x_35; lean_objec x_32 = l_List_appendTR___rarg(x_30, x_12); x_33 = lean_array_mk(x_32); x_34 = lean_array_size(x_33); -x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_34, x_5, x_33); +x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_34, x_5, x_33); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_35); x_36 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3094,7 +3094,7 @@ x_69 = l_List_appendTR___rarg(x_68, x_66); x_70 = l_List_appendTR___rarg(x_30, x_69); x_71 = lean_array_mk(x_70); x_72 = lean_array_size(x_71); -x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_72, x_5, x_71); +x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_72, x_5, x_71); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_73); x_74 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3138,7 +3138,7 @@ x_88 = l_List_appendTR___rarg(x_87, x_85); x_89 = l_List_appendTR___rarg(x_30, x_88); x_90 = lean_array_mk(x_89); x_91 = lean_array_size(x_90); -x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_91, x_5, x_90); +x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_91, x_5, x_90); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_92); x_93 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3198,7 +3198,7 @@ x_109 = l_List_appendTR___rarg(x_108, x_106); x_110 = l_List_appendTR___rarg(x_30, x_109); x_111 = lean_array_mk(x_110); x_112 = lean_array_size(x_111); -x_113 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_112, x_5, x_111); +x_113 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_112, x_5, x_111); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_113); x_114 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3295,7 +3295,7 @@ x_141 = l_List_appendTR___rarg(x_140, x_138); x_142 = l_List_appendTR___rarg(x_30, x_141); x_143 = lean_array_mk(x_142); x_144 = lean_array_size(x_143); -x_145 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_144, x_5, x_143); +x_145 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_144, x_5, x_143); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_145); x_146 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3407,7 +3407,7 @@ x_175 = l_List_appendTR___rarg(x_174, x_172); x_176 = l_List_appendTR___rarg(x_30, x_175); x_177 = lean_array_mk(x_176); x_178 = lean_array_size(x_177); -x_179 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_178, x_5, x_177); +x_179 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_178, x_5, x_177); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_179); x_180 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3548,7 +3548,7 @@ x_217 = l_List_appendTR___rarg(x_216, x_214); x_218 = l_List_appendTR___rarg(x_30, x_217); x_219 = lean_array_mk(x_218); x_220 = lean_array_size(x_219); -x_221 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_220, x_5, x_219); +x_221 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_220, x_5, x_219); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_221); x_222 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3595,7 +3595,7 @@ lean_object* x_235; lean_object* x_236; size_t x_237; lean_object* x_238; lean_o x_235 = l_List_appendTR___rarg(x_233, x_12); x_236 = lean_array_mk(x_235); x_237 = lean_array_size(x_236); -x_238 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_237, x_5, x_236); +x_238 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_237, x_5, x_236); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_238); x_239 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3746,7 +3746,7 @@ x_277 = l_List_appendTR___rarg(x_276, x_274); x_278 = l_List_appendTR___rarg(x_233, x_277); x_279 = lean_array_mk(x_278); x_280 = lean_array_size(x_279); -x_281 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_280, x_5, x_279); +x_281 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_280, x_5, x_279); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_281); x_282 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3808,7 +3808,7 @@ lean_object* x_297; lean_object* x_298; size_t x_299; lean_object* x_300; lean_o x_297 = l_List_appendTR___rarg(x_295, x_12); x_298 = lean_array_mk(x_297); x_299 = lean_array_size(x_298); -x_300 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_299, x_5, x_298); +x_300 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_299, x_5, x_298); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_300); x_301 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -3959,7 +3959,7 @@ x_339 = l_List_appendTR___rarg(x_338, x_336); x_340 = l_List_appendTR___rarg(x_295, x_339); x_341 = lean_array_mk(x_340); x_342 = lean_array_size(x_341); -x_343 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_342, x_5, x_341); +x_343 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_342, x_5, x_341); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_343); x_344 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -4043,7 +4043,7 @@ lean_object* x_364; lean_object* x_365; size_t x_366; lean_object* x_367; lean_o x_364 = l_List_appendTR___rarg(x_362, x_12); x_365 = lean_array_mk(x_364); x_366 = lean_array_size(x_365); -x_367 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_366, x_5, x_365); +x_367 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_366, x_5, x_365); x_368 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_368, 0, x_367); x_369 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -4194,7 +4194,7 @@ x_407 = l_List_appendTR___rarg(x_406, x_404); x_408 = l_List_appendTR___rarg(x_362, x_407); x_409 = lean_array_mk(x_408); x_410 = lean_array_size(x_409); -x_411 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_410, x_5, x_409); +x_411 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_410, x_5, x_409); x_412 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_412, 0, x_411); x_413 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -4310,7 +4310,7 @@ x_22 = lean_unsigned_to_nat(0u); x_23 = lean_array_uset(x_3, x_2, x_22); x_24 = lean_array_size(x_21); x_25 = 0; -x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_24, x_25, x_21); +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_24, x_25, x_21); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); lean_dec(x_26); @@ -6344,7 +6344,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRefInfo(lean_object* x_1) { lean_object* x_2; lean_object* x_3; x_2 = l_Lean_Lsp_instToJsonRefInfo___closed__2; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -6651,7 +6651,7 @@ x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = lean_array_mk(x_6); x_10 = lean_array_size(x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_1, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_1, x_9); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = 1; @@ -6800,7 +6800,7 @@ lean_object* x_43; lean_object* x_44; size_t x_45; lean_object* x_46; lean_objec x_43 = l_List_appendTR___rarg(x_41, x_21); x_44 = lean_array_mk(x_43); x_45 = lean_array_size(x_44); -x_46 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_45, x_15, x_44); +x_46 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_45, x_15, x_44); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_46); x_47 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -6906,7 +6906,7 @@ x_83 = l_List_appendTR___rarg(x_82, x_80); x_84 = l_List_appendTR___rarg(x_41, x_83); x_85 = lean_array_mk(x_84); x_86 = lean_array_size(x_85); -x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_86, x_15, x_85); +x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_86, x_15, x_85); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_87); x_88 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -6958,7 +6958,7 @@ x_105 = l_List_appendTR___rarg(x_104, x_102); x_106 = l_List_appendTR___rarg(x_41, x_105); x_107 = lean_array_mk(x_106); x_108 = lean_array_size(x_107); -x_109 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_108, x_15, x_107); +x_109 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_108, x_15, x_107); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_109); x_110 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7026,7 +7026,7 @@ x_129 = l_List_appendTR___rarg(x_128, x_126); x_130 = l_List_appendTR___rarg(x_41, x_129); x_131 = lean_array_mk(x_130); x_132 = lean_array_size(x_131); -x_133 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_132, x_15, x_131); +x_133 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_132, x_15, x_131); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_133); x_134 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7131,7 +7131,7 @@ x_164 = l_List_appendTR___rarg(x_163, x_161); x_165 = l_List_appendTR___rarg(x_41, x_164); x_166 = lean_array_mk(x_165); x_167 = lean_array_size(x_166); -x_168 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_167, x_15, x_166); +x_168 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_167, x_15, x_166); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_168); x_169 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7251,7 +7251,7 @@ x_201 = l_List_appendTR___rarg(x_200, x_198); x_202 = l_List_appendTR___rarg(x_41, x_201); x_203 = lean_array_mk(x_202); x_204 = lean_array_size(x_203); -x_205 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_204, x_15, x_203); +x_205 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_204, x_15, x_203); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_205); x_206 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7400,7 +7400,7 @@ x_246 = l_List_appendTR___rarg(x_245, x_243); x_247 = l_List_appendTR___rarg(x_41, x_246); x_248 = lean_array_mk(x_247); x_249 = lean_array_size(x_248); -x_250 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_249, x_15, x_248); +x_250 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_249, x_15, x_248); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_250); x_251 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7455,7 +7455,7 @@ lean_object* x_267; lean_object* x_268; size_t x_269; lean_object* x_270; lean_o x_267 = l_List_appendTR___rarg(x_265, x_21); x_268 = lean_array_mk(x_267); x_269 = lean_array_size(x_268); -x_270 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_269, x_15, x_268); +x_270 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_269, x_15, x_268); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_270); x_271 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7614,7 +7614,7 @@ x_312 = l_List_appendTR___rarg(x_311, x_309); x_313 = l_List_appendTR___rarg(x_265, x_312); x_314 = lean_array_mk(x_313); x_315 = lean_array_size(x_314); -x_316 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_315, x_15, x_314); +x_316 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_315, x_15, x_314); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_316); x_317 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7684,7 +7684,7 @@ lean_object* x_335; lean_object* x_336; size_t x_337; lean_object* x_338; lean_o x_335 = l_List_appendTR___rarg(x_333, x_21); x_336 = lean_array_mk(x_335); x_337 = lean_array_size(x_336); -x_338 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_337, x_15, x_336); +x_338 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_337, x_15, x_336); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_338); x_339 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7843,7 +7843,7 @@ x_380 = l_List_appendTR___rarg(x_379, x_377); x_381 = l_List_appendTR___rarg(x_333, x_380); x_382 = lean_array_mk(x_381); x_383 = lean_array_size(x_382); -x_384 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_383, x_15, x_382); +x_384 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_383, x_15, x_382); lean_ctor_set_tag(x_12, 4); lean_ctor_set(x_12, 0, x_384); x_385 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -7935,7 +7935,7 @@ lean_object* x_408; lean_object* x_409; size_t x_410; lean_object* x_411; lean_o x_408 = l_List_appendTR___rarg(x_406, x_21); x_409 = lean_array_mk(x_408); x_410 = lean_array_size(x_409); -x_411 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_410, x_15, x_409); +x_411 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_410, x_15, x_409); x_412 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_412, 0, x_411); x_413 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -8094,7 +8094,7 @@ x_454 = l_List_appendTR___rarg(x_453, x_451); x_455 = l_List_appendTR___rarg(x_406, x_454); x_456 = lean_array_mk(x_455); x_457 = lean_array_size(x_456); -x_458 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_457, x_15, x_456); +x_458 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_457, x_15, x_456); x_459 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_459, 0, x_458); x_460 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -8241,7 +8241,7 @@ lean_object* x_506; lean_object* x_507; size_t x_508; lean_object* x_509; lean_o x_506 = l_List_appendTR___rarg(x_504, x_482); x_507 = lean_array_mk(x_506); x_508 = lean_array_size(x_507); -x_509 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_508, x_475, x_507); +x_509 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_508, x_475, x_507); if (lean_is_scalar(x_490)) { x_510 = lean_alloc_ctor(4, 1, 0); } else { @@ -8405,7 +8405,7 @@ x_552 = l_List_appendTR___rarg(x_551, x_549); x_553 = l_List_appendTR___rarg(x_504, x_552); x_554 = lean_array_mk(x_553); x_555 = lean_array_size(x_554); -x_556 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_555, x_475, x_554); +x_556 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_555, x_475, x_554); if (lean_is_scalar(x_490)) { x_557 = lean_alloc_ctor(4, 1, 0); } else { @@ -8573,7 +8573,7 @@ lean_object* x_607; lean_object* x_608; size_t x_609; lean_object* x_610; lean_o x_607 = l_List_appendTR___rarg(x_605, x_582); x_608 = lean_array_mk(x_607); x_609 = lean_array_size(x_608); -x_610 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_609, x_575, x_608); +x_610 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_609, x_575, x_608); if (lean_is_scalar(x_591)) { x_611 = lean_alloc_ctor(4, 1, 0); } else { @@ -8737,7 +8737,7 @@ x_653 = l_List_appendTR___rarg(x_652, x_650); x_654 = l_List_appendTR___rarg(x_605, x_653); x_655 = lean_array_mk(x_654); x_656 = lean_array_size(x_655); -x_657 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_656, x_575, x_655); +x_657 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_656, x_575, x_655); if (lean_is_scalar(x_591)) { x_658 = lean_alloc_ctor(4, 1, 0); } else { @@ -9210,7 +9210,7 @@ lean_inc(x_23); lean_dec(x_19); x_100 = l_Lean_Lsp_instToJsonRefInfo___closed__2; lean_inc(x_6); -x_101 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(x_6, x_100); +x_101 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__2(x_6, x_100); if (lean_obj_tag(x_101) == 0) { uint8_t x_102; @@ -10101,7 +10101,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1923____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10251,7 +10251,7 @@ x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = lean_array_mk(x_6); x_10 = lean_array_size(x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_1, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_1, x_9); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = 1; @@ -10375,7 +10375,7 @@ lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_objec x_44 = l_List_appendTR___rarg(x_42, x_39); x_45 = lean_array_mk(x_44); x_46 = lean_array_size(x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_46, x_16, x_45); +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_46, x_16, x_45); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_47); x_48 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10481,7 +10481,7 @@ x_84 = l_List_appendTR___rarg(x_83, x_81); x_85 = l_List_appendTR___rarg(x_42, x_84); x_86 = lean_array_mk(x_85); x_87 = lean_array_size(x_86); -x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_87, x_16, x_86); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_87, x_16, x_86); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_88); x_89 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10533,7 +10533,7 @@ x_106 = l_List_appendTR___rarg(x_105, x_103); x_107 = l_List_appendTR___rarg(x_42, x_106); x_108 = lean_array_mk(x_107); x_109 = lean_array_size(x_108); -x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_109, x_16, x_108); +x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_109, x_16, x_108); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_110); x_111 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10601,7 +10601,7 @@ x_130 = l_List_appendTR___rarg(x_129, x_127); x_131 = l_List_appendTR___rarg(x_42, x_130); x_132 = lean_array_mk(x_131); x_133 = lean_array_size(x_132); -x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_133, x_16, x_132); +x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_133, x_16, x_132); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_134); x_135 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10706,7 +10706,7 @@ x_165 = l_List_appendTR___rarg(x_164, x_162); x_166 = l_List_appendTR___rarg(x_42, x_165); x_167 = lean_array_mk(x_166); x_168 = lean_array_size(x_167); -x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_168, x_16, x_167); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_168, x_16, x_167); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_169); x_170 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10826,7 +10826,7 @@ x_202 = l_List_appendTR___rarg(x_201, x_199); x_203 = l_List_appendTR___rarg(x_42, x_202); x_204 = lean_array_mk(x_203); x_205 = lean_array_size(x_204); -x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_205, x_16, x_204); +x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_205, x_16, x_204); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_206); x_207 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -10975,7 +10975,7 @@ x_247 = l_List_appendTR___rarg(x_246, x_244); x_248 = l_List_appendTR___rarg(x_42, x_247); x_249 = lean_array_mk(x_248); x_250 = lean_array_size(x_249); -x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_250, x_16, x_249); +x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_250, x_16, x_249); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_251); x_252 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11031,7 +11031,7 @@ lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_o x_269 = l_List_appendTR___rarg(x_267, x_263); x_270 = lean_array_mk(x_269); x_271 = lean_array_size(x_270); -x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_271, x_16, x_270); +x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_271, x_16, x_270); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_272); x_273 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11190,7 +11190,7 @@ x_314 = l_List_appendTR___rarg(x_313, x_311); x_315 = l_List_appendTR___rarg(x_267, x_314); x_316 = lean_array_mk(x_315); x_317 = lean_array_size(x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_317, x_16, x_316); +x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_317, x_16, x_316); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_318); x_319 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11261,7 +11261,7 @@ lean_object* x_338; lean_object* x_339; size_t x_340; lean_object* x_341; lean_o x_338 = l_List_appendTR___rarg(x_336, x_331); x_339 = lean_array_mk(x_338); x_340 = lean_array_size(x_339); -x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_340, x_16, x_339); +x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_340, x_16, x_339); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_341); x_342 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11420,7 +11420,7 @@ x_383 = l_List_appendTR___rarg(x_382, x_380); x_384 = l_List_appendTR___rarg(x_336, x_383); x_385 = lean_array_mk(x_384); x_386 = lean_array_size(x_385); -x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_386, x_16, x_385); +x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_386, x_16, x_385); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_387); x_388 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11513,7 +11513,7 @@ lean_object* x_412; lean_object* x_413; size_t x_414; lean_object* x_415; lean_o x_412 = l_List_appendTR___rarg(x_410, x_405); x_413 = lean_array_mk(x_412); x_414 = lean_array_size(x_413); -x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_414, x_16, x_413); +x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_414, x_16, x_413); x_416 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11672,7 +11672,7 @@ x_458 = l_List_appendTR___rarg(x_457, x_455); x_459 = l_List_appendTR___rarg(x_410, x_458); x_460 = lean_array_mk(x_459); x_461 = lean_array_size(x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_461, x_16, x_460); +x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_461, x_16, x_460); x_463 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_463, 0, x_462); x_464 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -11820,7 +11820,7 @@ lean_object* x_510; lean_object* x_511; size_t x_512; lean_object* x_513; lean_o x_510 = l_List_appendTR___rarg(x_508, x_503); x_511 = lean_array_mk(x_510); x_512 = lean_array_size(x_511); -x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_512, x_479, x_511); +x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_512, x_479, x_511); if (lean_is_scalar(x_493)) { x_514 = lean_alloc_ctor(4, 1, 0); } else { @@ -11984,7 +11984,7 @@ x_556 = l_List_appendTR___rarg(x_555, x_553); x_557 = l_List_appendTR___rarg(x_508, x_556); x_558 = lean_array_mk(x_557); x_559 = lean_array_size(x_558); -x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_559, x_479, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_559, x_479, x_558); if (lean_is_scalar(x_493)) { x_561 = lean_alloc_ctor(4, 1, 0); } else { @@ -12153,7 +12153,7 @@ lean_object* x_611; lean_object* x_612; size_t x_613; lean_object* x_614; lean_o x_611 = l_List_appendTR___rarg(x_609, x_604); x_612 = lean_array_mk(x_611); x_613 = lean_array_size(x_612); -x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_613, x_579, x_612); +x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_613, x_579, x_612); if (lean_is_scalar(x_594)) { x_615 = lean_alloc_ctor(4, 1, 0); } else { @@ -12317,7 +12317,7 @@ x_657 = l_List_appendTR___rarg(x_656, x_654); x_658 = l_List_appendTR___rarg(x_609, x_657); x_659 = lean_array_mk(x_658); x_660 = lean_array_size(x_659); -x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_660, x_579, x_659); +x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_660, x_579, x_659); if (lean_is_scalar(x_594)) { x_662 = lean_alloc_ctor(4, 1, 0); } else { @@ -12460,7 +12460,7 @@ lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_objec x_44 = l_List_appendTR___rarg(x_42, x_39); x_45 = lean_array_mk(x_44); x_46 = lean_array_size(x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_46, x_16, x_45); +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_46, x_16, x_45); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_47); x_48 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12566,7 +12566,7 @@ x_84 = l_List_appendTR___rarg(x_83, x_81); x_85 = l_List_appendTR___rarg(x_42, x_84); x_86 = lean_array_mk(x_85); x_87 = lean_array_size(x_86); -x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_87, x_16, x_86); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_87, x_16, x_86); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_88); x_89 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12618,7 +12618,7 @@ x_106 = l_List_appendTR___rarg(x_105, x_103); x_107 = l_List_appendTR___rarg(x_42, x_106); x_108 = lean_array_mk(x_107); x_109 = lean_array_size(x_108); -x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_109, x_16, x_108); +x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_109, x_16, x_108); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_110); x_111 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12686,7 +12686,7 @@ x_130 = l_List_appendTR___rarg(x_129, x_127); x_131 = l_List_appendTR___rarg(x_42, x_130); x_132 = lean_array_mk(x_131); x_133 = lean_array_size(x_132); -x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_133, x_16, x_132); +x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_133, x_16, x_132); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_134); x_135 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12791,7 +12791,7 @@ x_165 = l_List_appendTR___rarg(x_164, x_162); x_166 = l_List_appendTR___rarg(x_42, x_165); x_167 = lean_array_mk(x_166); x_168 = lean_array_size(x_167); -x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_168, x_16, x_167); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_168, x_16, x_167); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_169); x_170 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -12911,7 +12911,7 @@ x_202 = l_List_appendTR___rarg(x_201, x_199); x_203 = l_List_appendTR___rarg(x_42, x_202); x_204 = lean_array_mk(x_203); x_205 = lean_array_size(x_204); -x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_205, x_16, x_204); +x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_205, x_16, x_204); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_206); x_207 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13060,7 +13060,7 @@ x_247 = l_List_appendTR___rarg(x_246, x_244); x_248 = l_List_appendTR___rarg(x_42, x_247); x_249 = lean_array_mk(x_248); x_250 = lean_array_size(x_249); -x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_250, x_16, x_249); +x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_250, x_16, x_249); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_251); x_252 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13116,7 +13116,7 @@ lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_o x_269 = l_List_appendTR___rarg(x_267, x_263); x_270 = lean_array_mk(x_269); x_271 = lean_array_size(x_270); -x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_271, x_16, x_270); +x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_271, x_16, x_270); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_272); x_273 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13275,7 +13275,7 @@ x_314 = l_List_appendTR___rarg(x_313, x_311); x_315 = l_List_appendTR___rarg(x_267, x_314); x_316 = lean_array_mk(x_315); x_317 = lean_array_size(x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_317, x_16, x_316); +x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_317, x_16, x_316); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_318); x_319 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13346,7 +13346,7 @@ lean_object* x_338; lean_object* x_339; size_t x_340; lean_object* x_341; lean_o x_338 = l_List_appendTR___rarg(x_336, x_331); x_339 = lean_array_mk(x_338); x_340 = lean_array_size(x_339); -x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_340, x_16, x_339); +x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_340, x_16, x_339); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_341); x_342 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13505,7 +13505,7 @@ x_383 = l_List_appendTR___rarg(x_382, x_380); x_384 = l_List_appendTR___rarg(x_336, x_383); x_385 = lean_array_mk(x_384); x_386 = lean_array_size(x_385); -x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_386, x_16, x_385); +x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_386, x_16, x_385); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_387); x_388 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13598,7 +13598,7 @@ lean_object* x_412; lean_object* x_413; size_t x_414; lean_object* x_415; lean_o x_412 = l_List_appendTR___rarg(x_410, x_405); x_413 = lean_array_mk(x_412); x_414 = lean_array_size(x_413); -x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_414, x_16, x_413); +x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_414, x_16, x_413); x_416 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13757,7 +13757,7 @@ x_458 = l_List_appendTR___rarg(x_457, x_455); x_459 = l_List_appendTR___rarg(x_410, x_458); x_460 = lean_array_mk(x_459); x_461 = lean_array_size(x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_461, x_16, x_460); +x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_461, x_16, x_460); x_463 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_463, 0, x_462); x_464 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -13905,7 +13905,7 @@ lean_object* x_510; lean_object* x_511; size_t x_512; lean_object* x_513; lean_o x_510 = l_List_appendTR___rarg(x_508, x_503); x_511 = lean_array_mk(x_510); x_512 = lean_array_size(x_511); -x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_512, x_479, x_511); +x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_512, x_479, x_511); if (lean_is_scalar(x_493)) { x_514 = lean_alloc_ctor(4, 1, 0); } else { @@ -14069,7 +14069,7 @@ x_556 = l_List_appendTR___rarg(x_555, x_553); x_557 = l_List_appendTR___rarg(x_508, x_556); x_558 = lean_array_mk(x_557); x_559 = lean_array_size(x_558); -x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_559, x_479, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_559, x_479, x_558); if (lean_is_scalar(x_493)) { x_561 = lean_alloc_ctor(4, 1, 0); } else { @@ -14238,7 +14238,7 @@ lean_object* x_611; lean_object* x_612; size_t x_613; lean_object* x_614; lean_o x_611 = l_List_appendTR___rarg(x_609, x_604); x_612 = lean_array_mk(x_611); x_613 = lean_array_size(x_612); -x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_613, x_579, x_612); +x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_613, x_579, x_612); if (lean_is_scalar(x_594)) { x_615 = lean_alloc_ctor(4, 1, 0); } else { @@ -14402,7 +14402,7 @@ x_657 = l_List_appendTR___rarg(x_656, x_654); x_658 = l_List_appendTR___rarg(x_609, x_657); x_659 = lean_array_mk(x_658); x_660 = lean_array_size(x_659); -x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_660, x_579, x_659); +x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_660, x_579, x_659); if (lean_is_scalar(x_594)) { x_662 = lean_alloc_ctor(4, 1, 0); } else { @@ -14545,7 +14545,7 @@ lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_objec x_44 = l_List_appendTR___rarg(x_42, x_39); x_45 = lean_array_mk(x_44); x_46 = lean_array_size(x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_46, x_16, x_45); +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_46, x_16, x_45); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_47); x_48 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14651,7 +14651,7 @@ x_84 = l_List_appendTR___rarg(x_83, x_81); x_85 = l_List_appendTR___rarg(x_42, x_84); x_86 = lean_array_mk(x_85); x_87 = lean_array_size(x_86); -x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_87, x_16, x_86); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_87, x_16, x_86); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_88); x_89 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14703,7 +14703,7 @@ x_106 = l_List_appendTR___rarg(x_105, x_103); x_107 = l_List_appendTR___rarg(x_42, x_106); x_108 = lean_array_mk(x_107); x_109 = lean_array_size(x_108); -x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_109, x_16, x_108); +x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_109, x_16, x_108); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_110); x_111 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14771,7 +14771,7 @@ x_130 = l_List_appendTR___rarg(x_129, x_127); x_131 = l_List_appendTR___rarg(x_42, x_130); x_132 = lean_array_mk(x_131); x_133 = lean_array_size(x_132); -x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_133, x_16, x_132); +x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_133, x_16, x_132); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_134); x_135 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14876,7 +14876,7 @@ x_165 = l_List_appendTR___rarg(x_164, x_162); x_166 = l_List_appendTR___rarg(x_42, x_165); x_167 = lean_array_mk(x_166); x_168 = lean_array_size(x_167); -x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_168, x_16, x_167); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_168, x_16, x_167); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_169); x_170 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -14996,7 +14996,7 @@ x_202 = l_List_appendTR___rarg(x_201, x_199); x_203 = l_List_appendTR___rarg(x_42, x_202); x_204 = lean_array_mk(x_203); x_205 = lean_array_size(x_204); -x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_205, x_16, x_204); +x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_205, x_16, x_204); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_206); x_207 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15145,7 +15145,7 @@ x_247 = l_List_appendTR___rarg(x_246, x_244); x_248 = l_List_appendTR___rarg(x_42, x_247); x_249 = lean_array_mk(x_248); x_250 = lean_array_size(x_249); -x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_250, x_16, x_249); +x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_250, x_16, x_249); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_251); x_252 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15201,7 +15201,7 @@ lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_o x_269 = l_List_appendTR___rarg(x_267, x_263); x_270 = lean_array_mk(x_269); x_271 = lean_array_size(x_270); -x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_271, x_16, x_270); +x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_271, x_16, x_270); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_272); x_273 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15360,7 +15360,7 @@ x_314 = l_List_appendTR___rarg(x_313, x_311); x_315 = l_List_appendTR___rarg(x_267, x_314); x_316 = lean_array_mk(x_315); x_317 = lean_array_size(x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_317, x_16, x_316); +x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_317, x_16, x_316); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_318); x_319 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15431,7 +15431,7 @@ lean_object* x_338; lean_object* x_339; size_t x_340; lean_object* x_341; lean_o x_338 = l_List_appendTR___rarg(x_336, x_331); x_339 = lean_array_mk(x_338); x_340 = lean_array_size(x_339); -x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_340, x_16, x_339); +x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_340, x_16, x_339); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_341); x_342 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15590,7 +15590,7 @@ x_383 = l_List_appendTR___rarg(x_382, x_380); x_384 = l_List_appendTR___rarg(x_336, x_383); x_385 = lean_array_mk(x_384); x_386 = lean_array_size(x_385); -x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_386, x_16, x_385); +x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_386, x_16, x_385); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_387); x_388 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15683,7 +15683,7 @@ lean_object* x_412; lean_object* x_413; size_t x_414; lean_object* x_415; lean_o x_412 = l_List_appendTR___rarg(x_410, x_405); x_413 = lean_array_mk(x_412); x_414 = lean_array_size(x_413); -x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_414, x_16, x_413); +x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_414, x_16, x_413); x_416 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15842,7 +15842,7 @@ x_458 = l_List_appendTR___rarg(x_457, x_455); x_459 = l_List_appendTR___rarg(x_410, x_458); x_460 = lean_array_mk(x_459); x_461 = lean_array_size(x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_461, x_16, x_460); +x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_461, x_16, x_460); x_463 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_463, 0, x_462); x_464 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -15990,7 +15990,7 @@ lean_object* x_510; lean_object* x_511; size_t x_512; lean_object* x_513; lean_o x_510 = l_List_appendTR___rarg(x_508, x_503); x_511 = lean_array_mk(x_510); x_512 = lean_array_size(x_511); -x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_512, x_479, x_511); +x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_512, x_479, x_511); if (lean_is_scalar(x_493)) { x_514 = lean_alloc_ctor(4, 1, 0); } else { @@ -16154,7 +16154,7 @@ x_556 = l_List_appendTR___rarg(x_555, x_553); x_557 = l_List_appendTR___rarg(x_508, x_556); x_558 = lean_array_mk(x_557); x_559 = lean_array_size(x_558); -x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_559, x_479, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_559, x_479, x_558); if (lean_is_scalar(x_493)) { x_561 = lean_alloc_ctor(4, 1, 0); } else { @@ -16323,7 +16323,7 @@ lean_object* x_611; lean_object* x_612; size_t x_613; lean_object* x_614; lean_o x_611 = l_List_appendTR___rarg(x_609, x_604); x_612 = lean_array_mk(x_611); x_613 = lean_array_size(x_612); -x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_613, x_579, x_612); +x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_613, x_579, x_612); if (lean_is_scalar(x_594)) { x_615 = lean_alloc_ctor(4, 1, 0); } else { @@ -16487,7 +16487,7 @@ x_657 = l_List_appendTR___rarg(x_656, x_654); x_658 = l_List_appendTR___rarg(x_609, x_657); x_659 = lean_array_mk(x_658); x_660 = lean_array_size(x_659); -x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_660, x_579, x_659); +x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_660, x_579, x_659); if (lean_is_scalar(x_594)) { x_662 = lean_alloc_ctor(4, 1, 0); } else { @@ -16630,7 +16630,7 @@ lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_objec x_44 = l_List_appendTR___rarg(x_42, x_39); x_45 = lean_array_mk(x_44); x_46 = lean_array_size(x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_46, x_16, x_45); +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_46, x_16, x_45); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_47); x_48 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -16736,7 +16736,7 @@ x_84 = l_List_appendTR___rarg(x_83, x_81); x_85 = l_List_appendTR___rarg(x_42, x_84); x_86 = lean_array_mk(x_85); x_87 = lean_array_size(x_86); -x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_87, x_16, x_86); +x_88 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_87, x_16, x_86); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_88); x_89 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -16788,7 +16788,7 @@ x_106 = l_List_appendTR___rarg(x_105, x_103); x_107 = l_List_appendTR___rarg(x_42, x_106); x_108 = lean_array_mk(x_107); x_109 = lean_array_size(x_108); -x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_109, x_16, x_108); +x_110 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_109, x_16, x_108); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_110); x_111 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -16856,7 +16856,7 @@ x_130 = l_List_appendTR___rarg(x_129, x_127); x_131 = l_List_appendTR___rarg(x_42, x_130); x_132 = lean_array_mk(x_131); x_133 = lean_array_size(x_132); -x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_133, x_16, x_132); +x_134 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_133, x_16, x_132); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_134); x_135 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -16961,7 +16961,7 @@ x_165 = l_List_appendTR___rarg(x_164, x_162); x_166 = l_List_appendTR___rarg(x_42, x_165); x_167 = lean_array_mk(x_166); x_168 = lean_array_size(x_167); -x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_168, x_16, x_167); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_168, x_16, x_167); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_169); x_170 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17081,7 +17081,7 @@ x_202 = l_List_appendTR___rarg(x_201, x_199); x_203 = l_List_appendTR___rarg(x_42, x_202); x_204 = lean_array_mk(x_203); x_205 = lean_array_size(x_204); -x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_205, x_16, x_204); +x_206 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_205, x_16, x_204); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_206); x_207 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17230,7 +17230,7 @@ x_247 = l_List_appendTR___rarg(x_246, x_244); x_248 = l_List_appendTR___rarg(x_42, x_247); x_249 = lean_array_mk(x_248); x_250 = lean_array_size(x_249); -x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_250, x_16, x_249); +x_251 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_250, x_16, x_249); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_251); x_252 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17286,7 +17286,7 @@ lean_object* x_269; lean_object* x_270; size_t x_271; lean_object* x_272; lean_o x_269 = l_List_appendTR___rarg(x_267, x_263); x_270 = lean_array_mk(x_269); x_271 = lean_array_size(x_270); -x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_271, x_16, x_270); +x_272 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_271, x_16, x_270); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_272); x_273 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17445,7 +17445,7 @@ x_314 = l_List_appendTR___rarg(x_313, x_311); x_315 = l_List_appendTR___rarg(x_267, x_314); x_316 = lean_array_mk(x_315); x_317 = lean_array_size(x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_317, x_16, x_316); +x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_317, x_16, x_316); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_318); x_319 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17516,7 +17516,7 @@ lean_object* x_338; lean_object* x_339; size_t x_340; lean_object* x_341; lean_o x_338 = l_List_appendTR___rarg(x_336, x_331); x_339 = lean_array_mk(x_338); x_340 = lean_array_size(x_339); -x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_340, x_16, x_339); +x_341 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_340, x_16, x_339); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_341); x_342 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17675,7 +17675,7 @@ x_383 = l_List_appendTR___rarg(x_382, x_380); x_384 = l_List_appendTR___rarg(x_336, x_383); x_385 = lean_array_mk(x_384); x_386 = lean_array_size(x_385); -x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_386, x_16, x_385); +x_387 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_386, x_16, x_385); lean_ctor_set_tag(x_13, 4); lean_ctor_set(x_13, 0, x_387); x_388 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17768,7 +17768,7 @@ lean_object* x_412; lean_object* x_413; size_t x_414; lean_object* x_415; lean_o x_412 = l_List_appendTR___rarg(x_410, x_405); x_413 = lean_array_mk(x_412); x_414 = lean_array_size(x_413); -x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_414, x_16, x_413); +x_415 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_414, x_16, x_413); x_416 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -17927,7 +17927,7 @@ x_458 = l_List_appendTR___rarg(x_457, x_455); x_459 = l_List_appendTR___rarg(x_410, x_458); x_460 = lean_array_mk(x_459); x_461 = lean_array_size(x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_461, x_16, x_460); +x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_461, x_16, x_460); x_463 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_463, 0, x_462); x_464 = l_Lean_Lsp_instToJsonRefInfo___closed__2; @@ -18075,7 +18075,7 @@ lean_object* x_510; lean_object* x_511; size_t x_512; lean_object* x_513; lean_o x_510 = l_List_appendTR___rarg(x_508, x_503); x_511 = lean_array_mk(x_510); x_512 = lean_array_size(x_511); -x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_512, x_479, x_511); +x_513 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_512, x_479, x_511); if (lean_is_scalar(x_493)) { x_514 = lean_alloc_ctor(4, 1, 0); } else { @@ -18239,7 +18239,7 @@ x_556 = l_List_appendTR___rarg(x_555, x_553); x_557 = l_List_appendTR___rarg(x_508, x_556); x_558 = lean_array_mk(x_557); x_559 = lean_array_size(x_558); -x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_559, x_479, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_559, x_479, x_558); if (lean_is_scalar(x_493)) { x_561 = lean_alloc_ctor(4, 1, 0); } else { @@ -18408,7 +18408,7 @@ lean_object* x_611; lean_object* x_612; size_t x_613; lean_object* x_614; lean_o x_611 = l_List_appendTR___rarg(x_609, x_604); x_612 = lean_array_mk(x_611); x_613 = lean_array_size(x_612); -x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_613, x_579, x_612); +x_614 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_613, x_579, x_612); if (lean_is_scalar(x_594)) { x_615 = lean_alloc_ctor(4, 1, 0); } else { @@ -18572,7 +18572,7 @@ x_657 = l_List_appendTR___rarg(x_656, x_654); x_658 = l_List_appendTR___rarg(x_609, x_657); x_659 = lean_array_mk(x_658); x_660 = lean_array_size(x_659); -x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_660, x_579, x_659); +x_661 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_660, x_579, x_659); if (lean_is_scalar(x_594)) { x_662 = lean_alloc_ctor(4, 1, 0); } else { @@ -18654,7 +18654,7 @@ x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_8); lean_ctor_set(x_21, 1, x_20); x_22 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_21, x_22); +x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_21, x_22); x_24 = l_Lean_Json_mkObj(x_23); return x_24; } @@ -18682,7 +18682,7 @@ x_34 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_34, 0, x_8); lean_ctor_set(x_34, 1, x_33); x_35 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_36 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_34, x_35); +x_36 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_34, x_35); x_37 = l_Lean_Json_mkObj(x_36); return x_37; } @@ -18713,7 +18713,7 @@ x_46 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_46, 0, x_8); lean_ctor_set(x_46, 1, x_45); x_47 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_48 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_46, x_47); +x_48 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_46, x_47); x_49 = l_Lean_Json_mkObj(x_48); return x_49; } @@ -18741,7 +18741,7 @@ x_59 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_59, 0, x_8); lean_ctor_set(x_59, 1, x_58); x_60 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_61 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_59, x_60); +x_61 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_59, x_60); x_62 = l_Lean_Json_mkObj(x_61); return x_62; } @@ -19144,7 +19144,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); x_11 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } @@ -19271,7 +19271,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJson { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanStaleDependencyParams____x40_Lean_Data_Lsp_Internal___hyg_2211____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -19355,7 +19355,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_RefInfo_toJsonParentDecl____x40_Lean_Data_Lsp_Internal___hyg_723____closed__4; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } diff --git a/stage0/stdlib/Lean/Data/Lsp/Ipc.c b/stage0/stdlib/Lean/Data/Lsp/Ipc.c index 6cffa9fe10ec..77aeb0e48fea 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Ipc.c +++ b/stage0/stdlib/Lean/Data/Lsp/Ipc.c @@ -48,6 +48,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_Ipc_collectDiagnostics_loop___boxed(lean_obj static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__25; static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__52; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Lsp_Ipc_shutdown___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__20; static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Lsp_Ipc_shutdown___spec__5___closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_Ipc_shutdown(lean_object*, lean_object*, lean_object*); @@ -68,7 +69,6 @@ static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__2; LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Lsp_Ipc_collectDiagnostics___spec__3(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspRequest___at_Lean_Lsp_Ipc_collectDiagnostics___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Lsp_Ipc_shutdown___spec__2(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_Ipc_readResponseAs___closed__31; lean_object* lean_nat_to_int(lean_object*); lean_object* l_IO_FS_Stream_readLspMessage(lean_object*, lean_object*); @@ -2613,7 +2613,7 @@ x_155 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_155, 0, x_153); lean_ctor_set(x_155, 1, x_154); x_156 = l_Lean_Lsp_Ipc_readResponseAs___closed__8; -x_157 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_156, x_150); +x_157 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_156, x_150); lean_dec(x_150); switch (lean_obj_tag(x_147)) { case 0: @@ -3130,7 +3130,7 @@ x_281 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_281, 0, x_279); lean_ctor_set(x_281, 1, x_280); x_282 = l_Lean_Lsp_Ipc_readResponseAs___closed__8; -x_283 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_282, x_276); +x_283 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1455____spec__7(x_282, x_276); lean_dec(x_276); switch (lean_obj_tag(x_273)) { case 0: diff --git a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c index 44b2494a499b..5ab65ad2fb1c 100644 --- a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c +++ b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c @@ -81,7 +81,6 @@ static lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCall___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____spec__3___rarg(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCall; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__7; @@ -118,7 +117,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption_ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4438____closed__5; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8921____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlight___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__28; @@ -136,6 +134,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion___rarg(uint8_t, LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5646____spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238____lambda__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__110; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__15(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__5; @@ -212,8 +211,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__136; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__147; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238_(lean_object*); -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9856_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__12; @@ -222,9 +219,9 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238____lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11217____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__34; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__22; @@ -244,6 +241,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__122; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__23(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__8; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1(lean_object*, lean_object*); @@ -276,6 +274,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___rarg___boxed(lean_ob static lean_object* l_Lean_Lsp_instToJsonDocumentSymbol___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__9___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2783____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10869____closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__5; @@ -283,7 +282,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__15; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__43; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___closed__1; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__78; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3973____closed__7; static lean_object* l_Lean_Lsp_instFromJsonReferenceContext___closed__1; @@ -305,6 +303,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolTag___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4040_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__16; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__3; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__152; static lean_object* l_Lean_Lsp_instFromJsonHover___closed__1; @@ -326,7 +325,6 @@ static lean_object* l_Lean_Lsp_instToJsonCompletionItemTag___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__22(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__13; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8921_(uint8_t, uint8_t); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonHover; @@ -391,7 +389,6 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__40; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__34; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6337_(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__46; static lean_object* l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__13; @@ -409,7 +406,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqSemanticTokenType; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7312____closed__6; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____lambda__1(lean_object*); @@ -449,6 +445,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____lambda__1___boxed(lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__3; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__2; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__18; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__3; @@ -505,6 +502,7 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__15; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__6; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__84; static lean_object* l_Lean_Lsp_instFromJsonCompletionItem___closed__1; @@ -520,7 +518,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__15___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238____lambda__1___closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__5; -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__11; @@ -566,7 +563,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4098____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5646_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__16(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__32; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__45; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); @@ -617,7 +613,6 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__23; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__39; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4098____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__31; -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10159____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__86; @@ -657,7 +652,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l_Lean_Lsp_instBEqCallHierarchyItem___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__23___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__9; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolKind___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10284____closed__1; @@ -683,8 +677,8 @@ static lean_object* l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1; static lean_object* l_Lean_Lsp_instHashableSymbolTag___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__17; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__8; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__24; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__44; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4098____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); @@ -756,7 +750,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJso LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__21; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10994____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3(size_t, size_t, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -834,6 +827,7 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1851____rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2667____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__4; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__101; @@ -858,11 +852,13 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__70; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__130; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8939____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCompletionItem; LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2497____spec__5(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3175____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10081_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3803____closed__8; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__15; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__11; @@ -882,12 +878,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__10; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__18; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__46; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4438____closed__4; static lean_object* l_Lean_Lsp_instInhabitedCompletionItem___closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonReferenceContext___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__31; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRange; @@ -911,6 +909,7 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__37; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion(lean_object*); +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11217____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__3; @@ -932,12 +931,12 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprC LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__125; +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__2; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10800____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonRenameOptions___closed__1; LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2497____spec__2(lean_object*, lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__21; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__6; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1; @@ -996,7 +995,6 @@ LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hash static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__26; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__27; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__17; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10625_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionParams; @@ -1069,6 +1067,7 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7745____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__18; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__131; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10159____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -1114,7 +1113,6 @@ static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__ LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_toCtorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__23; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__12___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__21___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableSymbolKind; static lean_object* l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1; @@ -1167,10 +1165,10 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739_(lean_object*); static lean_object* l_Lean_Lsp_instBEqCompletionItem___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__14; LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2497____spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_toCtorIdx___boxed(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____closed__17; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____spec__2(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11217____closed__8; @@ -1184,12 +1182,13 @@ static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2783____closed__12; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__45; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9238____lambda__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2497_(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10465____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7188____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1(lean_object*); @@ -1232,7 +1231,6 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7312____closed__3; LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedCompletionItemKind; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__14; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1851____rarg___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4658_(lean_object*); @@ -1286,15 +1284,15 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion(lean_object*, lean_obj static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__148; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__15; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7438____closed__9; -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6337____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7920____closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__92; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__104; static lean_object* l_Lean_Lsp_instFromJsonRenameParams___closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__107; static lean_object* l_Lean_Lsp_instToJsonRenameParams___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonCompletionList___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__10; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__15___closed__1; @@ -1309,7 +1307,6 @@ static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__10; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__13___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__60; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10994____closed__12; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4658____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -1345,6 +1342,7 @@ static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind(uint8_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__12; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__8; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__17; @@ -1362,6 +1360,7 @@ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJ static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10284____closed__4; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____spec__3(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8085____lambda__11___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__132; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9750____closed__4; @@ -1394,6 +1393,7 @@ static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__44; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__66; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqInsertReplaceEdit; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5646____spec__2___rarg(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1___closed__1; @@ -1439,8 +1439,11 @@ static lean_object* l_Lean_Lsp_instFromJsonDefinitionParams___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10284____closed__5; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__10; +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokenType; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1___boxed(lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9157____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__102; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__140; @@ -1450,9 +1453,7 @@ static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__7; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3463____closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4438____closed__2; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_349____closed__135; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____closed__15; @@ -2070,7 +2071,7 @@ x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__22; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -2310,7 +2311,7 @@ x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_4); lean_ctor_set(x_16, 1, x_15); x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } @@ -5555,7 +5556,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5593,7 +5594,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__10; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -5631,7 +5632,7 @@ x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__15; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -5733,7 +5734,7 @@ x_8 = lean_box(0); x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -5741,7 +5742,7 @@ lean_ctor_set(x_12, 1, x_10); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_8); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1403____closed__15; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -5759,7 +5760,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_9); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -5800,7 +5801,7 @@ return x_10; else { uint8_t x_11; -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_7); +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_7); if (x_11 == 0) { uint8_t x_12; @@ -5810,7 +5811,7 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_5, x_8); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_5, x_8); return x_13; } } @@ -5853,9 +5854,9 @@ x_4 = lean_ctor_get(x_1, 2); x_5 = 0; x_6 = lean_string_hash(x_2); x_7 = lean_uint64_mix_hash(x_5, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_3); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_3); x_9 = lean_uint64_mix_hash(x_7, x_8); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_4); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_4); x_11 = lean_uint64_mix_hash(x_9, x_10); return x_11; } @@ -6239,7 +6240,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -6290,7 +6291,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -7509,7 +7510,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7547,7 +7548,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -7709,7 +7710,7 @@ lean_inc(x_56); lean_dec(x_47); x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__34; lean_inc(x_1); -x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_57); +x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_57); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -7928,7 +7929,7 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_2, 0); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -8031,7 +8032,33 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -8055,7 +8082,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -8075,7 +8102,7 @@ lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; le x_5 = lean_ctor_get(x_2, 0); x_6 = lean_array_size(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(x_6, x_7, x_5); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_8); x_9 = lean_alloc_ctor(0, 2, 0); @@ -8095,7 +8122,7 @@ lean_inc(x_12); lean_dec(x_2); x_13 = lean_array_size(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(x_13, x_14, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(x_13, x_14, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = lean_alloc_ctor(0, 2, 0); @@ -8142,7 +8169,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; -x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_15, x_3); +x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_15, x_3); x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__16; x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__1(x_17, x_4); lean_dec(x_4); @@ -8151,12 +8178,12 @@ x_20 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__28; x_22 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__3(x_21, x_6); x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__34; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_7); +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_7); x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__40; -x_26 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_25, x_8); +x_26 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(x_25, x_8); lean_dec(x_8); x_27 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__46; -x_28 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(x_27, x_9); +x_28 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(x_27, x_9); x_29 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_29, 0, x_28); lean_ctor_set(x_29, 1, x_13); @@ -8182,7 +8209,7 @@ x_36 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_36, 0, x_14); lean_ctor_set(x_36, 1, x_35); x_37 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_38 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_36, x_37); +x_38 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_36, x_37); x_39 = l_Lean_Json_mkObj(x_38); return x_39; } @@ -8196,7 +8223,16 @@ lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -8204,7 +8240,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__5(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__6(x_4, x_5, x_3); return x_6; } } @@ -8326,7 +8362,7 @@ else lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_2, 0); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6317_(x_6, x_7); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6132_(x_6, x_7); return x_8; } } @@ -8915,7 +8951,7 @@ lean_object* x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; x_70 = lean_ctor_get(x_4, 0); lean_inc(x_70); lean_dec(x_4); -x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_70); +x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_70); lean_dec(x_70); x_72 = 13; x_73 = lean_uint64_mix_hash(x_71, x_72); @@ -9410,7 +9446,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2783____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -9611,7 +9647,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_8); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -9788,7 +9824,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -9825,7 +9861,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -9908,7 +9944,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -9920,7 +9956,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -9935,7 +9971,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -9978,7 +10014,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6159_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5974_(x_2); lean_dec(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); @@ -9992,7 +10028,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_9, x_8); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155____spec__1(x_9, x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_6); @@ -10000,7 +10036,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -10026,7 +10062,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6211_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6026_(x_3); return x_4; } } @@ -10203,7 +10239,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1225____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -10375,7 +10411,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10412,7 +10448,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -10495,7 +10531,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -10507,7 +10543,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -10522,7 +10558,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -10629,7 +10665,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10666,7 +10702,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -10749,7 +10785,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -10761,7 +10797,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -10776,7 +10812,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -10883,7 +10919,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10920,7 +10956,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -11003,7 +11039,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -11015,7 +11051,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -11030,7 +11066,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -11137,7 +11173,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -11174,7 +11210,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -11257,7 +11293,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -11269,7 +11305,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -11284,7 +11320,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -11399,7 +11435,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3973____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -11483,7 +11519,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -11658,7 +11694,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -11696,7 +11732,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -11836,7 +11872,7 @@ x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_3); x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); @@ -11848,7 +11884,7 @@ lean_ctor_set(x_8, 1, x_7); x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_dec(x_2); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_9); x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -11879,7 +11915,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_8); lean_ctor_set(x_22, 1, x_21); x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_22, x_23); +x_24 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_22, x_23); x_25 = l_Lean_Json_mkObj(x_24); return x_25; } @@ -11994,7 +12030,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4321____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -12078,7 +12114,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -12185,7 +12221,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -12222,7 +12258,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -12305,7 +12341,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -12317,7 +12353,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -12332,7 +12368,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -12552,7 +12588,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -12574,7 +12610,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -12669,7 +12705,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -12739,7 +12775,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_1); x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -12752,7 +12788,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -21885,7 +21921,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_3; lean_object* x_4; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__1; lean_inc(x_2); -x_4 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_2, x_3); +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_2, x_3); if (lean_obj_tag(x_4) == 0) { uint8_t x_5; @@ -21924,7 +21960,7 @@ lean_inc(x_13); lean_dec(x_4); x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; lean_inc(x_2); -x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_2, x_14); +x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_2, x_14); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; @@ -22005,7 +22041,7 @@ lean_inc(x_35); lean_dec(x_26); x_36 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; lean_inc(x_2); -x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_2, x_36); +x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_2, x_36); if (lean_obj_tag(x_37) == 0) { uint8_t x_38; @@ -22047,7 +22083,7 @@ lean_inc(x_46); lean_dec(x_37); x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__20; lean_inc(x_2); -x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_2, x_47); +x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_2, x_47); if (lean_obj_tag(x_48) == 0) { uint8_t x_49; @@ -22351,8 +22387,8 @@ x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_14, x_4); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_6); +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_14, x_4); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_6); x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); @@ -22360,7 +22396,7 @@ lean_ctor_set(x_18, 1, x_16); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_12); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_7); +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_7); x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__20; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); @@ -22583,7 +22619,7 @@ x_35 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_35, 0, x_13); lean_ctor_set(x_35, 1, x_34); x_36 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_37 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_35, x_36); +x_37 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_35, x_36); x_38 = l_Lean_Json_mkObj(x_37); return x_38; } @@ -22733,8 +22769,8 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; -x_14 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_13, x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_5); +x_14 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_13, x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_5); x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); @@ -22742,7 +22778,7 @@ lean_ctor_set(x_17, 1, x_15); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_11); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_6); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_6); x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__20; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); @@ -22965,7 +23001,7 @@ x_34 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_34, 0, x_12); lean_ctor_set(x_34, 1, x_33); x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_36 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_34, x_35); +x_36 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_34, x_35); x_37 = l_Lean_Json_mkObj(x_36); return x_37; } @@ -23795,7 +23831,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947_(x_3); return x_4; } } @@ -24032,7 +24068,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24189,7 +24225,7 @@ x_45 = lean_ctor_get(x_36, 0); lean_inc(x_45); lean_dec(x_36); x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____closed__18; -x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_46); +x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_46); if (lean_obj_tag(x_47) == 0) { uint8_t x_48; @@ -24372,7 +24408,7 @@ lean_ctor_set(x_17, 1, x_15); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_10); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(x_5); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(x_5); x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____closed__13; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); @@ -24381,7 +24417,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_10); x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6018____closed__18; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_6); +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_23, x_6); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_10); @@ -24592,7 +24628,7 @@ x_33 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_33, 0, x_11); lean_ctor_set(x_33, 1, x_32); x_34 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_33, x_34); +x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_33, x_34); x_36 = l_Lean_Json_mkObj(x_35); return x_36; } @@ -24712,7 +24748,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24749,7 +24785,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -24832,7 +24868,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -24844,7 +24880,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -24859,7 +24895,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -25290,7 +25326,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -25407,7 +25443,7 @@ lean_inc(x_34); lean_dec(x_25); x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; lean_inc(x_1); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -25448,7 +25484,7 @@ lean_inc(x_45); lean_dec(x_36); x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__13; lean_inc(x_1); -x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_46); +x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_46); if (lean_obj_tag(x_47) == 0) { uint8_t x_48; @@ -25490,7 +25526,7 @@ lean_inc(x_56); lean_dec(x_47); x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; lean_inc(x_1); -x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_57); +x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_57); if (lean_obj_tag(x_58) == 0) { uint8_t x_59; @@ -25533,7 +25569,7 @@ lean_inc(x_67); lean_dec(x_58); x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__20; lean_inc(x_1); -x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_68); +x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_68); if (lean_obj_tag(x_69) == 0) { uint8_t x_70; @@ -25739,7 +25775,7 @@ lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__46; x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6912____spec__1(x_15, x_4); x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__10; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_17, x_5); +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_17, x_5); x_19 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_19, 0, x_6); x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6572____closed__13; @@ -25749,7 +25785,7 @@ lean_ctor_set(x_21, 1, x_19); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_13); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_7); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_7); x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_25 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_25, 0, x_24); @@ -25757,7 +25793,7 @@ lean_ctor_set(x_25, 1, x_23); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_13); -x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_8); +x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_8); x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5367____rarg___closed__20; x_29 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_29, 0, x_28); @@ -25766,7 +25802,7 @@ x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_13); x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2074____closed__40; -x_32 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_31, x_9); +x_32 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2414____spec__4(x_31, x_9); lean_dec(x_9); x_33 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_33, 0, x_32); @@ -25987,7 +26023,7 @@ x_44 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_44, 0, x_14); lean_ctor_set(x_44, 1, x_43); x_45 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_46 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_44, x_45); +x_46 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_44, x_45); x_47 = l_Lean_Json_mkObj(x_46); return x_47; } @@ -26157,7 +26193,7 @@ return x_28; else { uint8_t x_29; -x_29 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_8, x_16); +x_29 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_8, x_16); if (x_29 == 0) { uint8_t x_30; @@ -26167,7 +26203,7 @@ return x_30; else { uint8_t x_31; -x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_9, x_17); +x_31 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_9, x_17); if (x_31 == 0) { uint8_t x_32; @@ -26278,8 +26314,8 @@ x_12 = lean_uint64_mix_hash(x_10, x_11); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4846_(x_3); x_14 = lean_uint64_mix_hash(x_12, x_13); x_15 = lean_string_hash(x_6); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_7); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_8); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_7); +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_8); if (lean_obj_tag(x_4) == 0) { uint64_t x_46; @@ -26678,7 +26714,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -26717,7 +26753,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_6 = lean_array_uget(x_3, x_2); x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_3, x_2, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_6); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -27131,7 +27167,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -27178,7 +27214,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_7); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -27388,7 +27424,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -27675,7 +27711,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_7); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -32198,7 +32234,7 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_10); lean_ctor_set(x_19, 1, x_18); x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_21 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_19, x_20); +x_21 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_19, x_20); x_22 = l_Lean_Json_mkObj(x_21); return x_22; } @@ -32430,7 +32466,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -32468,7 +32504,7 @@ x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9936____closed__12; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -32612,7 +32648,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_7); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -32698,7 +32734,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -32768,7 +32804,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_1); x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -32781,7 +32817,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -32888,7 +32924,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -32925,7 +32961,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -33008,7 +33044,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -33020,7 +33056,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3133____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -33035,7 +33071,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -33338,7 +33374,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10465____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -33507,7 +33543,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; size_t x x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10465____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); lean_dec(x_1); @@ -33531,7 +33567,7 @@ x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_4); lean_ctor_set(x_15, 1, x_14); x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_15, x_16); +x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_15, x_16); x_18 = l_Lean_Json_mkObj(x_17); return x_18; } @@ -33629,7 +33665,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -33699,7 +33735,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_1); x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -33712,7 +33748,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -33985,7 +34021,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_8); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -34109,7 +34145,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10869____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -34193,7 +34229,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -34359,7 +34395,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -34397,7 +34433,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -34435,7 +34471,7 @@ x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10994____closed__9; -x_25 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -34528,7 +34564,7 @@ x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_3); x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); @@ -34540,7 +34576,7 @@ lean_ctor_set(x_8, 1, x_7); x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_dec(x_2); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_9); x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -34570,7 +34606,7 @@ x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_8); lean_ctor_set(x_21, 1, x_20); x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_21, x_22); +x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_21, x_22); x_24 = l_Lean_Json_mkObj(x_23); return x_24; } @@ -34677,7 +34713,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -34714,7 +34750,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -34797,7 +34833,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -34809,7 +34845,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2953____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -34824,7 +34860,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_7); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } diff --git a/stage0/stdlib/Lean/Data/Lsp/TextSync.c b/stage0/stdlib/Lean/Data/Lsp/TextSync.c index ee876cde5ce2..6e0902ce108b 100644 --- a/stage0/stdlib/Lean/Data/Lsp/TextSync.c +++ b/stage0/stdlib/Lean/Data/Lsp/TextSync.c @@ -24,6 +24,7 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCl static lean_object* l_Lean_Lsp_instToJsonSaveOptions___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentContentChangeEvent_hasToJson(lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDidOpenTextDocumentParams; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253_(lean_object*); @@ -36,10 +37,8 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpen LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_872____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_toCtorIdx(uint8_t); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__9; -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__7; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____spec__1___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__7; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent(lean_object*); @@ -47,27 +46,27 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveO static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__3; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__3; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__4; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSaveOptions; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__7; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164_(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__10; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____spec__1___closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__5; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__6; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__9; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__22; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__7; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__11; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__10; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__18; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__6; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__6; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_533____closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__6; @@ -81,6 +80,7 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSa static lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind___closed__1; static lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncKind___closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__11; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidOpenTextDocumentParams; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -89,13 +89,14 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_noConfusion(lean_object static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__4; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind___closed__2; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__2; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidChangeTextDocumentParams; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__11; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncKind___closed__4; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__11; @@ -103,7 +104,6 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOp LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_toCtorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__14; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__9; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__4; static lean_object* l_Lean_Lsp_instToJsonDidSaveTextDocumentParams___closed__1; @@ -111,13 +111,12 @@ static lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind___closed__3; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_872_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__20; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_989_(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_533_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_712_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__10; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____closed__5; @@ -132,9 +131,12 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextD static lean_object* l_Lean_Lsp_instFromJsonTextDocumentSyncKind___closed__4; static lean_object* l_Lean_Lsp_TextDocumentSyncKind_noConfusion___rarg___closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__12; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__9; static lean_object* l_Lean_Lsp_instFromJsonDidSaveTextDocumentParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__1; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__10; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidCloseTextDocumentParams; @@ -144,7 +146,6 @@ static lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncKind___closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncOptions; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_533____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_910____closed__2; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__15; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__14; @@ -169,7 +170,6 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextD static lean_object* l_Lean_Lsp_instFromJsonDidCloseTextDocumentParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSaveOptions; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentChangeRegistrationOptions; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__2; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__8; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__6; @@ -182,12 +182,13 @@ static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextD static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1223____closed__25; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__4; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754_(lean_object*); lean_object* lean_array_mk(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__10; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__8; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__4; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__4; @@ -204,7 +205,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_TextDocumentSyncKind_noConfusion___rarg___bo static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidCloseTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_1027____closed__3; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__6; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDidChangeTextDocumentParams___closed__1; static lean_object* l_Lean_Lsp_instToJsonTextDocumentSyncKind___closed__5; @@ -221,11 +221,11 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_754____closed__5; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____closed__8; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____spec__1(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidSaveTextDocumentParams; +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_585____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_164____lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__5; @@ -543,7 +543,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDi _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4990_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4805_(x_1); x_3 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -556,7 +556,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -582,7 +582,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5070_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4885_(x_3); return x_4; } } @@ -1020,7 +1020,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentChangeRegistrationOptions____x40_Lean_Data_Lsp_TextSync___hyg_253____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5864____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5679____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1169,7 +1169,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent(l lean_object* x_2; lean_object* x_3; x_2 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -1180,7 +1180,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); lean_dec(x_5); x_6 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_7 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_6); +x_7 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_6); if (lean_obj_tag(x_7) == 0) { uint8_t x_8; @@ -1233,7 +1233,7 @@ else lean_object* x_15; lean_object* x_16; lean_dec(x_3); x_15 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_16 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_15); +x_16 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_15); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -1285,7 +1285,7 @@ x_24 = lean_ctor_get(x_3, 0); lean_inc(x_24); lean_dec(x_3); x_25 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_26 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_25); +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_25); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; @@ -1349,7 +1349,7 @@ if (x_2 == 0) lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_6 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; lean_ctor_set(x_1, 1, x_5); lean_ctor_set(x_1, 0, x_6); @@ -1377,7 +1377,7 @@ x_15 = lean_ctor_get(x_1, 1); lean_inc(x_15); lean_inc(x_14); lean_dec(x_1); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_14); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_14); x_17 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); @@ -1467,7 +1467,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_12); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_12); x_15 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; lean_ctor_set(x_6, 1, x_14); lean_ctor_set(x_6, 0, x_15); @@ -1498,7 +1498,7 @@ x_25 = lean_ctor_get(x_6, 1); lean_inc(x_25); lean_inc(x_24); lean_dec(x_6); -x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_24); +x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_24); x_27 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; x_28 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_28, 0, x_27); @@ -1585,7 +1585,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDi lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2440_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2255_(x_2); x_4 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -1616,7 +1616,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_7); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -1669,7 +1669,7 @@ x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_3, x_2, x_7); x_9 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__1; lean_inc(x_6); -x_10 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_6, x_9); +x_10 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__2(x_6, x_9); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; @@ -1680,7 +1680,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_10, 0); lean_dec(x_12); x_13 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_6, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_6, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -1723,7 +1723,7 @@ else lean_object* x_23; lean_object* x_24; lean_dec(x_10); x_23 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_24 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_6, x_23); +x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_6, x_23); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -1769,7 +1769,7 @@ x_34 = lean_ctor_get(x_10, 0); lean_inc(x_34); lean_dec(x_10); x_35 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_36 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_6, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_6, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -2017,7 +2017,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2660____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2475____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2158,7 +2158,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDi lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_2); x_4 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -2171,7 +2171,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_8); +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1800____spec__2(x_9, x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_6); @@ -2179,7 +2179,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -2315,7 +2315,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2352,7 +2352,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l_Lean_Lsp_instFromJsonTextDocumentContentChangeEvent___closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1860____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2455,7 +2455,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -2572,7 +2572,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonSaveOptions____x40_Lean_Data_Lsp_TextSync___hyg_872____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2642,7 +2642,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDi _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2313_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2128_(x_1); x_3 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -2655,7 +2655,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -2741,7 +2741,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson { lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3011,7 +3011,7 @@ x_27 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_27, 0, x_11); lean_ctor_set(x_27, 1, x_26); x_28 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_27, x_28); +x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_27, x_28); x_30 = l_Lean_Json_mkObj(x_29); return x_30; } @@ -3026,7 +3026,7 @@ x_33 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_33, 0, x_11); lean_ctor_set(x_33, 1, x_32); x_34 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_33, x_34); +x_35 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_33, x_34); x_36 = l_Lean_Json_mkObj(x_35); return x_36; } @@ -3041,7 +3041,7 @@ x_39 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_39, 0, x_11); lean_ctor_set(x_39, 1, x_38); x_40 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_126____closed__2; -x_41 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_39, x_40); +x_41 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_39, x_40); x_42 = l_Lean_Json_mkObj(x_41); return x_42; } @@ -3503,7 +3503,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJson lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -3580,7 +3580,7 @@ lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__2; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -3620,7 +3620,7 @@ lean_inc(x_34); lean_dec(x_25); x_35 = l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139____closed__3; lean_inc(x_1); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2853____spec__1(x_1, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2668____spec__1(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; diff --git a/stage0/stdlib/Lean/Data/Lsp/Utf16.c b/stage0/stdlib/Lean/Data/Lsp/Utf16.c index b1128d2ef78d..c43a16aa59cb 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Utf16.c +++ b/stage0/stdlib/Lean/Data/Lsp/Utf16.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Data.Lsp.Utf16 -// Imports: Init.Data.String Init.Data.Array Lean.Data.Lsp.Basic Lean.Data.Position Lean.DeclarationRange +// Imports: Init.Data.String Lean.Data.Lsp.Basic Lean.Data.Position Lean.DeclarationRange #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -561,7 +561,6 @@ return x_2; } } lean_object* initialize_Init_Data_String(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_Array(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Lsp_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_Position(uint8_t builtin, lean_object*); lean_object* initialize_Lean_DeclarationRange(uint8_t builtin, lean_object*); @@ -573,9 +572,6 @@ _G_initialized = true; res = initialize_Init_Data_String(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Array(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_Data_Lsp_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Data/Lsp/Workspace.c b/stage0/stdlib/Lean/Data/Lsp/Workspace.c index 0ec53b37a480..efcad0521383 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Workspace.c +++ b/stage0/stdlib/Lean/Data/Lsp/Workspace.c @@ -40,7 +40,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_FileSystemWatcher_delete; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____closed__15; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____closed__16; static lean_object* l_Lean_Lsp_instFromJsonFileChangeType___closed__1; -lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__1; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____closed__6; static lean_object* l_Lean_Lsp_instFromJsonWorkspaceFolder___closed__1; @@ -52,7 +51,6 @@ static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWork static lean_object* l_Lean_Lsp_instToJsonDidChangeWatchedFilesRegistrationOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_816____closed__2; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____lambda__1(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_883_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__9; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileEvent____x40_Lean_Data_Lsp_Workspace___hyg_645____closed__8; @@ -70,7 +68,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFileChangeType___boxed(lean_object LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonFileChangeType(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDidChangeWatchedFilesParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__8; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_816____spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25_(lean_object*); @@ -95,7 +92,6 @@ static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonFileEv static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____closed__5; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____closed__2; static lean_object* l_Lean_Lsp_instFromJsonFileChangeType___closed__5; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileEvent____x40_Lean_Data_Lsp_Workspace___hyg_645____closed__11; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____spec__1___boxed(lean_object*, lean_object*); @@ -117,6 +113,7 @@ static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFile LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_457____spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonFileEvent; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_390____closed__1; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonFileSystemWatcher___closed__1; static lean_object* l_Lean_Lsp_instToJsonFileChangeType___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_311_(lean_object*); @@ -155,8 +152,10 @@ static lean_object* l_Lean_Lsp_instFromJsonDidChangeWatchedFilesParams___closed_ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_816_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonFileEvent____x40_Lean_Data_Lsp_Workspace___hyg_751____closed__9; static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_77____closed__12; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FileChangeType_noConfusion(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FileChangeType_noConfusion___rarg(uint8_t, uint8_t, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileEvent____x40_Lean_Data_Lsp_Workspace___hyg_645____closed__1; @@ -176,6 +175,7 @@ static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonFileEv static lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonDidChangeWatchedFilesParams____x40_Lean_Data_Lsp_Workspace___hyg_816____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_FileSystemWatcher_create; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDidChangeWatchedFilesRegistrationOptions; static lean_object* l_Lean_Lsp_instToJsonFileChangeType___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_457____spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -242,7 +242,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); x_15 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -280,7 +280,7 @@ x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); x_31 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_30, x_31); +x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_30, x_31); x_33 = l_Lean_Json_mkObj(x_32); return x_33; } @@ -478,7 +478,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJso lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -515,7 +515,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -746,7 +746,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJso lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -783,7 +783,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJsonFileSystemWatcher____x40_Lean_Data_Lsp_Workspace___hyg_205____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2482____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2297____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -923,7 +923,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } @@ -952,7 +952,7 @@ x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_8); lean_ctor_set(x_21, 1, x_20); x_22 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_21, x_22); +x_23 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_21, x_22); x_24 = l_Lean_Json_mkObj(x_23); return x_24; } @@ -979,7 +979,7 @@ x_32 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_32, 0, x_8); lean_ctor_set(x_32, 1, x_31); x_33 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_34 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_32, x_33); +x_34 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_32, x_33); x_35 = l_Lean_Json_mkObj(x_34); return x_35; } @@ -1011,7 +1011,7 @@ x_44 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); x_45 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_46 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_44, x_45); +x_46 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_44, x_45); x_47 = l_Lean_Json_mkObj(x_46); return x_47; } @@ -1049,7 +1049,7 @@ x_56 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_56, 0, x_42); lean_ctor_set(x_56, 1, x_55); x_57 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_58 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_56, x_57); +x_58 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_56, x_57); x_59 = l_Lean_Json_mkObj(x_58); return x_59; } @@ -1477,7 +1477,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); x_11 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } @@ -2158,7 +2158,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_fromJso lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_947____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2419,7 +2419,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); x_11 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } @@ -2431,7 +2431,7 @@ x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_8); lean_ctor_set(x_15, 1, x_14); x_16 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_15, x_16); +x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_15, x_16); x_18 = l_Lean_Json_mkObj(x_17); return x_18; } @@ -2443,7 +2443,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_8); lean_ctor_set(x_20, 1, x_19); x_21 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -2841,7 +2841,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); x_11 = l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonWorkspaceFolder____x40_Lean_Data_Lsp_Workspace___hyg_25____closed__3; -x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_10, x_11); +x_12 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_10, x_11); x_13 = l_Lean_Json_mkObj(x_12); return x_13; } diff --git a/stage0/stdlib/Lean/Declaration.c b/stage0/stdlib/Lean/Declaration.c index 5a0a47012492..70a88d518b1d 100644 --- a/stage0/stdlib/Lean/Declaration.c +++ b/stage0/stdlib/Lean/Declaration.c @@ -20,9 +20,10 @@ LEAN_EXPORT lean_object* lean_mk_axiom_val(lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lean_ConstantInfo_value_x21___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkQuotValEx___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3020____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RecursorVal_isUnsafeEx___boxed(lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqDefinitionVal____x40_Lean_Declaration___hyg_864____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__3; LEAN_EXPORT lean_object* l_Lean_mkInductiveValEx___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqInductiveType____x40_Lean_Declaration___hyg_1590____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__19; @@ -32,9 +33,11 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedRecursorRule; LEAN_EXPORT lean_object* l_Lean_OpaqueVal_isUnsafeEx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_QuotVal_kindEx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_type(lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__5; static lean_object* l_Lean_instBEqAxiomVal___closed__1; LEAN_EXPORT uint8_t lean_constructor_val_is_unsafe(lean_object*); static lean_object* l_Lean_instInhabitedOpaqueVal___closed__1; +static lean_object* l_Lean_Declaration_getNames___closed__8; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_Declaration_forExprM___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_definitionVal_x21___boxed(lean_object*); @@ -43,6 +46,7 @@ LEAN_EXPORT lean_object* l_Lean_instBEqConstantVal; LEAN_EXPORT lean_object* l_Lean_mkAxiomValEx___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedRecursorVal___closed__1; static lean_object* l_Lean_instReprDefinitionSafety___closed__1; +LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_isUnsafeInductiveDeclEx___boxed(lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__6; LEAN_EXPORT lean_object* l_Lean_RecursorVal_getMajorInduct(lean_object*); @@ -54,7 +58,6 @@ LEAN_EXPORT lean_object* l_Lean_ConstantInfo_isTheorem___boxed(lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__2; LEAN_EXPORT lean_object* l_Lean_instInhabitedTheoremVal; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_value_x3f___boxed(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_reducibility_hints_regular(uint32_t); lean_object* l_List_foldlM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_quot_val(lean_object*, lean_object*, lean_object*, uint8_t); @@ -83,6 +86,7 @@ static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety___ LEAN_EXPORT uint8_t l_Lean_ConstantInfo_isUnsafe(lean_object*); static lean_object* l_Lean_ConstantInfo_inductiveVal_x21___closed__1; LEAN_EXPORT uint8_t lean_recursor_is_unsafe(lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__10; static lean_object* l_Lean_instInhabitedQuotVal___closed__1; LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__3(lean_object*); @@ -100,13 +104,14 @@ LEAN_EXPORT lean_object* l_Lean_ConstantInfo_toConstantVal___boxed(lean_object*) LEAN_EXPORT lean_object* l_Lean_ConstantInfo_value_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* lean_mk_inductive_decl(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_instBEqConstantVal___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_quot_val_kind(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RecursorVal_kEx___boxed(lean_object*); static lean_object* l_Lean_instInhabitedConstantVal___closed__1; LEAN_EXPORT lean_object* l_Lean_Declaration_forExprM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_definition_val_get_safety(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqDefinitionSafety; static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__8; LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_toCtorIdx___boxed(lean_object*); @@ -129,18 +134,18 @@ LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqDeclaration____x40_L LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_Declaration_forExprM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqInductiveType; static lean_object* l_Lean_ConstantInfo_value_x21___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_ConstantInfo_isCtor(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__2___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedConstantVal___closed__2; LEAN_EXPORT uint8_t l_Lean_InductiveVal_isNested(lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_definitionVal_x21(lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__2; +static lean_object* l_Lean_Declaration_getNames___closed__6; LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqAxiomVal____x40_Lean_Declaration___hyg_552_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_QuotKind_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Lean_InductiveVal_numCtors___boxed(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqInductiveType____x40_Lean_Declaration___hyg_1590____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_definition_val(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); @@ -164,6 +169,7 @@ LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lea LEAN_EXPORT lean_object* l_Lean_ConstantInfo_numLevelParams___boxed(lean_object*); lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Declaration_getNames(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqTheoremVal____x40_Lean_Declaration___hyg_1095____boxed(lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -179,14 +185,17 @@ LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__2( LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqDefinitionVal; LEAN_EXPORT uint8_t l_Lean_instInhabitedDefinitionSafety; +static lean_object* l_Lean_Declaration_getNames___closed__12; static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__10; LEAN_EXPORT lean_object* l_Lean_instBEqDeclaration; LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__1; LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ReducibilityHints_lt___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_recursor_k(lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqRecursorVal; LEAN_EXPORT lean_object* l_Lean_mkRecName(lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Declaration_getNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_all(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructor____x40_Lean_Declaration___hyg_1475____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_Declaration_forExprM___spec__1(lean_object*); @@ -216,6 +225,7 @@ LEAN_EXPORT lean_object* l_Lean_QuotKind_noConfusion___rarg(uint8_t, uint8_t, le uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedInductiveVal___closed__1; LEAN_EXPORT lean_object* l_Lean_instInhabitedConstantInfo; +static lean_object* l_Lean_Declaration_getNames___closed__11; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Declaration_definitionVal_x21___closed__2; LEAN_EXPORT uint8_t l_Lean_ConstantInfo_isInductive(lean_object*); @@ -227,18 +237,18 @@ static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety___ LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_name___boxed(lean_object*); static lean_object* l_Lean_DefinitionSafety_noConfusion___rarg___closed__1; +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__11; LEAN_EXPORT lean_object* lean_mk_recursor_val(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_ReducibilityHints_isAbbrev___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_hasValue___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqInductiveType____x40_Lean_Declaration___hyg_1590_(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3020_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqConstructorVal; LEAN_EXPORT lean_object* l_Lean_instReprDefinitionSafety; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_hints(lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__3; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqDefinitionSafety____x40_Lean_Declaration___hyg_666____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_ReducibilityHints_isAbbrev(lean_object*); @@ -247,9 +257,12 @@ LEAN_EXPORT lean_object* l_Lean_InductiveVal_isNested___boxed(lean_object*); static lean_object* l_Lean_instBEqTheoremVal___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l_Lean_instInhabitedInductiveType___closed__1; +LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg___lambda__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instBEqRecursorRule___closed__1; LEAN_EXPORT lean_object* l_Lean_mkDefinitionValEx___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Declaration_getNames___spec__2(lean_object*, lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Lean_instBEqConstructorVal___closed__1; static lean_object* l_Lean_ConstantInfo_inductiveVal_x21___closed__2; LEAN_EXPORT lean_object* l_Lean_ReducibilityHints_instOrd; @@ -259,6 +272,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_ConstantInfo_inductiveVal_x21___spec_ LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__1; LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructor____x40_Lean_Declaration___hyg_1475_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__20; static lean_object* l_Lean_instInhabitedDefinitionVal___closed__1; LEAN_EXPORT lean_object* l_Lean_instInhabitedOpaqueVal; @@ -274,8 +288,11 @@ LEAN_EXPORT lean_object* l_Lean_ConstantInfo_levelParams___boxed(lean_object*); static lean_object* l_Lean_Declaration_definitionVal_x21___closed__3; static lean_object* l_Lean_Declaration_definitionVal_x21___closed__4; lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__7; LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_noConfusion___rarg___lambda__1___boxed(lean_object*); static lean_object* l_Lean_instInhabitedRecursorRule___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3146____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Declaration_getNames___closed__4; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_isCtor___boxed(lean_object*); LEAN_EXPORT lean_object* lean_mk_inductive_val(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____boxed(lean_object*, lean_object*); @@ -284,7 +301,6 @@ LEAN_EXPORT lean_object* l_Lean_instBEqOpaqueVal; LEAN_EXPORT lean_object* l_Lean_instInhabitedAxiomVal; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__5; -LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____spec__1___boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqAxiomVal____x40_Lean_Declaration___hyg_552____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_ConstantInfo_isPartial(lean_object*); @@ -294,6 +310,7 @@ LEAN_EXPORT lean_object* l_Lean_ConstantInfo_isPartial___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_InductiveVal_numCtors(lean_object*); LEAN_EXPORT lean_object* l_Lean_QuotKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_684____closed__18; +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3146_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqDeclaration____x40_Lean_Declaration___hyg_1774____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqReducibilityHints____x40_Lean_Declaration___hyg_28_(lean_object*, lean_object*); static lean_object* l_Lean_instBEqDefinitionVal___closed__1; @@ -2508,6 +2525,285 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Declaration_getNames___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 0); +lean_dec(x_8); +x_9 = lean_ctor_get(x_5, 0); +lean_inc(x_9); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_9); +{ +lean_object* _tmp_0 = x_7; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +lean_dec(x_1); +x_12 = lean_ctor_get(x_5, 0); +lean_inc(x_12); +lean_dec(x_5); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_2); +x_1 = x_11; +x_2 = x_13; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Declaration_getNames___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Quot", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Declaration_getNames___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("mk", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Declaration_getNames___closed__1; +x_2 = l_Lean_Declaration_getNames___closed__3; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("lift", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Declaration_getNames___closed__1; +x_2 = l_Lean_Declaration_getNames___closed__5; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ind", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Declaration_getNames___closed__1; +x_2 = l_Lean_Declaration_getNames___closed__7; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Declaration_getNames___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Declaration_getNames___closed__6; +x_2 = l_Lean_Declaration_getNames___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Declaration_getNames___closed__4; +x_2 = l_Lean_Declaration_getNames___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Declaration_getNames___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Declaration_getNames___closed__2; +x_2 = l_Lean_Declaration_getNames___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Declaration_getNames(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 4: +{ +lean_object* x_2; +x_2 = l_Lean_Declaration_getNames___closed__12; +return x_2; +} +case 5: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_box(0); +x_5 = l_List_mapTR_loop___at_Lean_Declaration_getNames___spec__1(x_3, x_4); +return x_5; +} +case 6: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_1, 2); +lean_inc(x_6); +lean_dec(x_1); +x_7 = lean_box(0); +x_8 = l_List_mapTR_loop___at_Lean_Declaration_getNames___spec__2(x_6, x_7); +return x_8; +} +default: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +return x_13; +} +} +} +} LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -3230,7 +3526,7 @@ x_1 = l_Lean_instInhabitedConstructorVal___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; @@ -3321,11 +3617,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3336,7 +3632,7 @@ static lean_object* _init_l_Lean_instBEqConstructorVal___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900____boxed), 2, 0); return x_1; } } @@ -3416,7 +3712,7 @@ x_1 = l_Lean_instInhabitedRecursorRule___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3020_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3146_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -3452,11 +3748,11 @@ return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3020____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3146____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3020_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3146_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3467,7 +3763,7 @@ static lean_object* _init_l_Lean_instBEqRecursorRule___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3020____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3146____boxed), 2, 0); return x_1; } } @@ -3508,7 +3804,7 @@ x_1 = l_Lean_instInhabitedRecursorVal___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3541,7 +3837,7 @@ x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_1, 1); x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3020_(x_6, x_8); +x_10 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_3146_(x_6, x_8); if (x_10 == 0) { uint8_t x_11; @@ -3558,7 +3854,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; uint8_t x_21; uint8_t x_26; @@ -3640,7 +3936,7 @@ return x_37; else { uint8_t x_38; -x_38 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____spec__1(x_9, x_18); +x_38 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____spec__1(x_9, x_18); if (x_38 == 0) { uint8_t x_39; @@ -3723,22 +4019,22 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____spec__1(x_1, x_2); +x_3 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3749,7 +4045,7 @@ static lean_object* _init_l_Lean_instBEqRecursorVal___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332____boxed), 2, 0); return x_1; } } @@ -4492,7 +4788,7 @@ static lean_object* _init_l_Lean_ConstantInfo_value_x21___closed__3() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Declaration_definitionVal_x21___closed__1; x_2 = l_Lean_ConstantInfo_value_x21___closed__1; -x_3 = lean_unsigned_to_nat(458u); +x_3 = lean_unsigned_to_nat(471u); x_4 = lean_unsigned_to_nat(31u); x_5 = l_Lean_ConstantInfo_value_x21___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4505,7 +4801,7 @@ static lean_object* _init_l_Lean_ConstantInfo_value_x21___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_Declaration_definitionVal_x21___closed__1; x_2 = l_Lean_ConstantInfo_value_x21___closed__1; -x_3 = lean_unsigned_to_nat(457u); +x_3 = lean_unsigned_to_nat(470u); x_4 = lean_unsigned_to_nat(62u); x_5 = l_Lean_ConstantInfo_value_x21___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4711,7 +5007,7 @@ static lean_object* _init_l_Lean_ConstantInfo_inductiveVal_x21___closed__3() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Declaration_definitionVal_x21___closed__1; x_2 = l_Lean_ConstantInfo_inductiveVal_x21___closed__1; -x_3 = lean_unsigned_to_nat(478u); +x_3 = lean_unsigned_to_nat(491u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_ConstantInfo_inductiveVal_x21___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4965,6 +5261,30 @@ l_Lean_Declaration_definitionVal_x21___closed__3 = _init_l_Lean_Declaration_defi lean_mark_persistent(l_Lean_Declaration_definitionVal_x21___closed__3); l_Lean_Declaration_definitionVal_x21___closed__4 = _init_l_Lean_Declaration_definitionVal_x21___closed__4(); lean_mark_persistent(l_Lean_Declaration_definitionVal_x21___closed__4); +l_Lean_Declaration_getNames___closed__1 = _init_l_Lean_Declaration_getNames___closed__1(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__1); +l_Lean_Declaration_getNames___closed__2 = _init_l_Lean_Declaration_getNames___closed__2(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__2); +l_Lean_Declaration_getNames___closed__3 = _init_l_Lean_Declaration_getNames___closed__3(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__3); +l_Lean_Declaration_getNames___closed__4 = _init_l_Lean_Declaration_getNames___closed__4(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__4); +l_Lean_Declaration_getNames___closed__5 = _init_l_Lean_Declaration_getNames___closed__5(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__5); +l_Lean_Declaration_getNames___closed__6 = _init_l_Lean_Declaration_getNames___closed__6(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__6); +l_Lean_Declaration_getNames___closed__7 = _init_l_Lean_Declaration_getNames___closed__7(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__7); +l_Lean_Declaration_getNames___closed__8 = _init_l_Lean_Declaration_getNames___closed__8(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__8); +l_Lean_Declaration_getNames___closed__9 = _init_l_Lean_Declaration_getNames___closed__9(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__9); +l_Lean_Declaration_getNames___closed__10 = _init_l_Lean_Declaration_getNames___closed__10(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__10); +l_Lean_Declaration_getNames___closed__11 = _init_l_Lean_Declaration_getNames___closed__11(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__11); +l_Lean_Declaration_getNames___closed__12 = _init_l_Lean_Declaration_getNames___closed__12(); +lean_mark_persistent(l_Lean_Declaration_getNames___closed__12); l_Lean_instInhabitedInductiveVal___closed__1 = _init_l_Lean_instInhabitedInductiveVal___closed__1(); lean_mark_persistent(l_Lean_instInhabitedInductiveVal___closed__1); l_Lean_instInhabitedInductiveVal = _init_l_Lean_instInhabitedInductiveVal(); diff --git a/stage0/stdlib/Lean/Elab/App.c b/stage0/stdlib/Lean/Elab/App.c index 000d0ab12645..9e4e0df469c1 100644 --- a/stage0/stdlib/Lean/Elab/App.c +++ b/stage0/stdlib/Lean/Elab/App.c @@ -304,7 +304,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___clos 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*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4____closed__9; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_ElabElim_finalize___spec__8___boxed(lean_object**); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__17(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplicitArg___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 uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__15(lean_object*, lean_object*, size_t, size_t); @@ -1091,7 +1091,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__3; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__21(lean_object*, lean_object*, size_t, size_t); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_8237____lambda__2___closed__15; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___closed__9; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__18; @@ -21737,7 +21737,8 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_main_module(x_8); +x_9 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_10 = lean_ctor_get(x_2, 10); lean_inc(x_10); lean_dec(x_2); @@ -21756,7 +21757,8 @@ lean_dec(x_5); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -x_15 = lean_environment_main_module(x_14); +x_15 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_16 = lean_ctor_get(x_2, 10); lean_inc(x_16); lean_dec(x_2); @@ -33181,7 +33183,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabA _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -lean_inc(x_1); x_5 = l_Lean_isRec___at___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabAsElim___spec__1(x_1, x_2, x_3, x_4); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); @@ -33235,6 +33236,7 @@ lean_object* x_5; x_5 = l_Lean_isRec___at___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabAsElim___spec__1(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } @@ -44466,8 +44468,6 @@ x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); -lean_dec(x_1); x_18 = lean_ctor_get(x_15, 0); lean_inc(x_18); lean_dec(x_15); @@ -44481,7 +44481,7 @@ x_22 = lean_ctor_get(x_19, 1); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_find(x_23, x_17); +x_24 = l_Lean_Environment_find_x3f(x_23, x_17); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; @@ -45133,7 +45133,7 @@ lean_dec(x_19); x_157 = lean_ctor_get(x_155, 0); lean_inc(x_157); lean_dec(x_155); -x_158 = lean_environment_find(x_157, x_17); +x_158 = l_Lean_Environment_find_x3f(x_157, x_17); if (lean_obj_tag(x_158) == 0) { lean_object* x_159; lean_object* x_160; @@ -45412,7 +45412,6 @@ else { lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_dec(x_5); -lean_dec(x_1); x_211 = lean_ctor_get(x_14, 1); lean_inc(x_211); lean_dec(x_14); @@ -45755,6 +45754,7 @@ x_27 = lean_box(0); x_28 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__1(x_12, x_3, x_1, x_24, x_23, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_9); lean_dec(x_24); +lean_dec(x_12); return x_28; } else @@ -46260,6 +46260,7 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_4); +lean_dec(x_1); return x_14; } } @@ -52517,6 +52518,7 @@ lean_dec(x_42); x_45 = lean_ctor_get(x_39, 1); lean_inc(x_45); x_46 = l_Lean_isPrivateNameFromImportedModule(x_44, x_45); +lean_dec(x_44); if (x_46 == 0) { lean_object* x_47; lean_object* x_48; @@ -52597,6 +52599,7 @@ lean_dec(x_60); x_63 = lean_ctor_get(x_39, 1); lean_inc(x_63); x_64 = l_Lean_isPrivateNameFromImportedModule(x_62, x_63); +lean_dec(x_62); if (x_64 == 0) { lean_object* x_65; lean_object* x_66; @@ -52721,6 +52724,7 @@ lean_dec(x_87); x_91 = lean_ctor_get(x_85, 1); lean_inc(x_91); x_92 = l_Lean_isPrivateNameFromImportedModule(x_90, x_91); +lean_dec(x_90); if (x_92 == 0) { lean_object* x_93; lean_object* x_94; @@ -52922,6 +52926,7 @@ lean_dec(x_129); x_133 = lean_ctor_get(x_127, 1); lean_inc(x_133); x_134 = l_Lean_isPrivateNameFromImportedModule(x_132, x_133); +lean_dec(x_132); if (x_134 == 0) { lean_object* x_135; lean_object* x_136; diff --git a/stage0/stdlib/Lean/Elab/Attributes.c b/stage0/stdlib/Lean/Elab/Attributes.c index eef27b8678f9..8a64bd28fe11 100644 --- a/stage0/stdlib/Lean/Elab/Attributes.c +++ b/stage0/stdlib/Lean/Elab/Attributes.c @@ -131,7 +131,7 @@ LEAN_EXPORT uint8_t l_Lean_Elab_instToFormatAttribute___lambda__1(lean_object*); static lean_object* l_Lean_Elab_elabAttr___rarg___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs(lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_toAttributeKind___closed__4; static lean_object* l_Lean_Elab_instInhabitedAttribute___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_elabAttrs___spec__1___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1109,7 +1109,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__ _start: { lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_environment_main_module(x_1); +x_18 = l_Lean_Environment_mainModule(x_1); x_19 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_19, 0, x_2); lean_ctor_set(x_19, 1, x_18); @@ -1572,7 +1572,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__ _start: { lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_environment_main_module(x_1); +x_18 = l_Lean_Environment_mainModule(x_1); x_19 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_19, 0, x_2); lean_ctor_set(x_19, 1, x_18); @@ -2344,6 +2344,7 @@ lean_object* x_17 = _args[16]; { lean_object* x_18; x_18 = l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__1___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_1); return x_18; } } @@ -2396,6 +2397,7 @@ lean_object* x_17 = _args[16]; { lean_object* x_18; x_18 = l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__3___rarg___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_1); return x_18; } } diff --git a/stage0/stdlib/Lean/Elab/Binders.c b/stage0/stdlib/Lean/Elab/Binders.c index a941f21cc528..45764cb262b2 100644 --- a/stage0/stdlib/Lean/Elab/Binders.c +++ b/stage0/stdlib/Lean/Elab/Binders.c @@ -564,7 +564,7 @@ lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop(lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange__1___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___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*); @@ -853,7 +853,8 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_main_module(x_8); +x_9 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_10 = lean_ctor_get(x_2, 10); lean_inc(x_10); lean_dec(x_2); @@ -872,7 +873,8 @@ lean_dec(x_5); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -x_15 = lean_environment_main_module(x_14); +x_15 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_16 = lean_ctor_get(x_2, 10); lean_inc(x_16); lean_dec(x_2); @@ -2700,7 +2702,8 @@ lean_dec(x_10); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -x_14 = lean_environment_main_module(x_13); +x_14 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_15 = lean_ctor_get(x_7, 10); lean_inc(x_15); x_16 = l_Lean_addMacroScope(x_14, x_1, x_15); @@ -3062,7 +3065,8 @@ x_30 = lean_ctor_get(x_28, 0); x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); lean_dec(x_30); -x_32 = lean_environment_main_module(x_31); +x_32 = l_Lean_Environment_mainModule(x_31); +lean_dec(x_31); x_33 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8; x_34 = l_Lean_addMacroScope(x_32, x_33, x_27); x_35 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__7; @@ -3094,7 +3098,8 @@ lean_dec(x_28); x_45 = lean_ctor_get(x_43, 0); lean_inc(x_45); lean_dec(x_43); -x_46 = lean_environment_main_module(x_45); +x_46 = l_Lean_Environment_mainModule(x_45); +lean_dec(x_45); x_47 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8; x_48 = l_Lean_addMacroScope(x_46, x_47, x_27); x_49 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__7; @@ -3175,7 +3180,8 @@ x_70 = lean_ctor_get(x_68, 0); x_71 = lean_ctor_get(x_70, 0); lean_inc(x_71); lean_dec(x_70); -x_72 = lean_environment_main_module(x_71); +x_72 = l_Lean_Environment_mainModule(x_71); +lean_dec(x_71); x_73 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13; x_74 = l_Lean_addMacroScope(x_72, x_73, x_67); x_75 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__12; @@ -3205,7 +3211,8 @@ lean_dec(x_68); x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); lean_dec(x_82); -x_85 = lean_environment_main_module(x_84); +x_85 = l_Lean_Environment_mainModule(x_84); +lean_dec(x_84); x_86 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13; x_87 = l_Lean_addMacroScope(x_85, x_86, x_67); x_88 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__12; @@ -7113,7 +7120,8 @@ x_25 = lean_ctor_get(x_23, 0); x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); lean_dec(x_25); -x_27 = lean_environment_main_module(x_26); +x_27 = l_Lean_Environment_mainModule(x_26); +lean_dec(x_26); x_28 = lean_ctor_get(x_7, 10); lean_inc(x_28); lean_dec(x_7); @@ -7135,7 +7143,8 @@ lean_dec(x_23); x_35 = lean_ctor_get(x_33, 0); lean_inc(x_35); lean_dec(x_33); -x_36 = lean_environment_main_module(x_35); +x_36 = l_Lean_Environment_mainModule(x_35); +lean_dec(x_35); x_37 = lean_ctor_get(x_7, 10); lean_inc(x_37); lean_dec(x_7); @@ -21937,7 +21946,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_13); +x_31 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_32 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_32, 0, x_25); lean_ctor_set(x_32, 1, x_31); @@ -22115,7 +22125,8 @@ lean_dec(x_26); x_80 = lean_ctor_get(x_78, 1); lean_inc(x_80); lean_dec(x_78); -x_81 = lean_environment_main_module(x_13); +x_81 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_82 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_82, 0, x_25); lean_ctor_set(x_82, 1, x_81); @@ -23163,7 +23174,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -23345,7 +23357,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); diff --git a/stage0/stdlib/Lean/Elab/BuiltinCommand.c b/stage0/stdlib/Lean/Elab/BuiltinCommand.c index f3709a66ad21..d44266250b2a 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -87,7 +87,6 @@ static lean_object* l_Lean_Elab_Command_elabOpen___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd_declRange__1___closed__2; static lean_object* l_Lean_resolveNamespaceCore___at_Lean_Elab_Command_elabExport___spec__4___closed__1; uint8_t l___private_Lean_Data_KVMap_0__Lean_beqDataValue____x40_Lean_Data_KVMap___hyg_70_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_hasNoErrorMessages(lean_object*); LEAN_EXPORT lean_object* l_Lean_popScope___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_popScopes___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -145,6 +144,7 @@ lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabImport___rarg___closed__1; static lean_object* l_Lean_Elab_Command_elabOmit___lambda__1___closed__1; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOptions___spec__1___lambda__2___closed__5; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3; lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM___at_Lean_Elab_Command_elabUniverse___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); @@ -158,6 +158,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabOpen___spe lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1___closed__2; lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_containsId___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInitQuot___boxed(lean_object*); @@ -244,7 +245,7 @@ static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed_ lean_object* l_Lean_ScopedEnvExtension_pushScope___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabVariable_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabExport___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_addUnivLevel___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getRefPos___at_Lean_Elab_Command_elabExport___spec__21___boxed(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInCmd__1(lean_object*); @@ -481,6 +482,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabVersion___rarg(lean_object*, le static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__5; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOmit___spec__13___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabVersion__1___closed__3; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabNonComputableSection__1___closed__2; @@ -562,7 +564,6 @@ LEAN_EXPORT lean_object* l_Lean_ensureNoOverload___at_Lean_Elab_Command_elabExpo lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__33; -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInclude___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__27; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabUniverse_declRange__1(lean_object*); @@ -628,6 +629,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabImport_declRange__1___c LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__8___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabExport___closed__4; static lean_object* l_Lean_Elab_Command_elabEnd___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_elabAddDeclDoc___spec__5___lambda__1___closed__1; uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_List_filterMapTR_go___at_Lean_resolveNamespace___spec__1(lean_object*, lean_object*); @@ -677,7 +679,6 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabExport___s static lean_object* l___regBuiltin_Lean_Elab_Command_elabChoice_declRange__1___closed__7; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_pushScope___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSection_declRange__1___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___closed__5; extern uint8_t l_System_Platform_isEmscripten; @@ -730,8 +731,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabReduce(lean_object*, lean_objec static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd_declRange__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabVariable_declRange__1___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__20___closed__2; -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabImport___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInCmd___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabOpen(lean_object*, lean_object*, lean_object*, lean_object*); @@ -754,7 +755,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Comma lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInCmd_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_pushScope___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3; static lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOpen___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabExit___rarg(lean_object*, lean_object*, lean_object*); @@ -801,6 +801,7 @@ lean_object* l_Lean_Elab_Command_withFreshMacroScope___rarg(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_resolveNameUsingNamespaces___at_Lean_Elab_Command_elabExport___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_withNamespace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Elab_Command_elabOmit___spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInitQuot___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSection_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(size_t, size_t, lean_object*); @@ -876,6 +877,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabChoice_declRange__1___c LEAN_EXPORT lean_object* l_Lean_Elab_Command_failIfSucceeds(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOmit___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*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOptions___spec__1___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabUniverse_declRange__1___closed__4; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi__1(lean_object*); @@ -989,7 +991,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUn static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__20___closed__1; lean_object* l_List_toString___at_Lean_ensureNoOverload___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_activateScoped___at_Lean_Elab_Command_elabOpen___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheck_declRange__1___closed__2; static lean_object* l_Lean_Elab_Command_elabOpen___closed__1; @@ -1049,6 +1050,7 @@ lean_object* lean_io_error_to_string(lean_object*); static lean_object* l_Lean_Elab_Command_elabSynth___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabImport_declRange__1___closed__2; static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___closed__1; static lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__4___closed__2; @@ -1064,6 +1066,7 @@ lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); static lean_object* l_Lean_resolveNamespaceCore___at_Lean_Elab_Command_elabExport___spec__4___closed__3; static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabWhere_describeOptions___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabOpen__1___closed__1; static uint8_t l_Lean_Elab_Command_elabVersion___rarg___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabInclude___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1074,7 +1077,6 @@ static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_ela LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabOpen__1(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); uint8_t l_List_beq___at___private_Lean_Data_OpenDecl_0__Lean_beqOpenDecl____x40_Lean_Data_OpenDecl___hyg_41____spec__1(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce_declRange__1(lean_object*); static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__37; static lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___closed__6; @@ -1104,6 +1106,7 @@ lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_obje uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___closed__2; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; +static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Command_elabExport___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); @@ -1140,7 +1143,6 @@ LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_Elab_Command_elabInclude static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__40; extern uint8_t l_System_Platform_isWindows; lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___spec__2(size_t, size_t, lean_object*); -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabVersion___rarg___lambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheck_declRange__1___closed__3; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabAddDeclDoc___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1183,7 +1185,6 @@ static lean_object* l_Lean_resolveNamespaceCore___at_Lean_Elab_Command_elabExpor LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOmit___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabOmit___spec__12(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__modelUsing____1___spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAddDeclDoc___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at_Lean_Elab_Command_elabExport___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5901,7 +5902,7 @@ lean_dec(x_22); x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec(x_23); -x_25 = l_Lean_KernelException_toMessageData(x_18, x_24); +x_25 = l_Lean_Kernel_Exception_toMessageData(x_18, x_24); x_26 = l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(x_25, x_1, x_2, x_21); return x_26; } @@ -8640,8 +8641,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -8682,8 +8682,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -11004,8 +11003,7 @@ x_9 = lean_ctor_get(x_6, 1); x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); -lean_inc(x_1); -x_11 = lean_environment_find(x_10, x_1); +x_11 = l_Lean_Environment_find_x3f(x_10, x_1); if (lean_obj_tag(x_11) == 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; @@ -11045,8 +11043,7 @@ lean_dec(x_6); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -lean_inc(x_1); -x_23 = lean_environment_find(x_22, x_1); +x_23 = l_Lean_Environment_find_x3f(x_22, x_1); if (lean_obj_tag(x_23) == 0) { uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -31828,6 +31825,31 @@ return x_3; } else { +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_mk_syntax_ident(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); @@ -31841,7 +31863,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -31866,7 +31888,7 @@ goto _start; } } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1() { _start: { lean_object* x_1; @@ -31874,7 +31896,7 @@ x_1 = lean_mk_string_unchecked("hiding", 6, 6); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2() { _start: { lean_object* x_1; @@ -31882,7 +31904,7 @@ x_1 = lean_mk_string_unchecked("renaming", 8, 8); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3() { _start: { lean_object* x_1; @@ -31890,19 +31912,19 @@ x_1 = lean_mk_string_unchecked("openRenamingItem", 16, 16); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__1; x_2 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__2; x_3 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc__1___closed__3; -x_4 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3; +x_4 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5() { +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5() { _start: { lean_object* x_1; @@ -31910,7 +31932,7 @@ x_1 = lean_mk_string_unchecked("→", 3, 1); return x_1; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { if (lean_obj_tag(x_6) == 0) @@ -32014,7 +32036,7 @@ lean_ctor_set(x_122, 1, x_126); lean_ctor_set(x_122, 0, x_117); x_127 = lean_array_size(x_16); x_128 = 0; -x_129 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_127, x_128, x_16); +x_129 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_127, x_128, x_16); x_130 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_131 = l_Array_append___rarg(x_130, x_129); lean_dec(x_129); @@ -32051,7 +32073,7 @@ lean_ctor_set(x_142, 0, x_117); lean_ctor_set(x_142, 1, x_141); x_143 = lean_array_size(x_16); x_144 = 0; -x_145 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_143, x_144, x_16); +x_145 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_143, x_144, x_16); x_146 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_147 = l_Array_append___rarg(x_146, x_145); lean_dec(x_145); @@ -32105,7 +32127,7 @@ lean_ctor_set(x_161, 0, x_117); lean_ctor_set(x_161, 1, x_160); x_162 = lean_array_size(x_16); x_163 = 0; -x_164 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_162, x_163, x_16); +x_164 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_162, x_163, x_16); x_165 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_166 = l_Array_append___rarg(x_165, x_164); lean_dec(x_164); @@ -32204,7 +32226,7 @@ lean_ctor_set_tag(x_34, 2); lean_ctor_set(x_34, 1, x_38); lean_ctor_set(x_34, 0, x_29); x_39 = lean_mk_syntax_ident(x_18); -x_40 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1; +x_40 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1; lean_inc(x_29); lean_ctor_set_tag(x_30, 2); lean_ctor_set(x_30, 1, x_40); @@ -32212,9 +32234,9 @@ lean_ctor_set(x_30, 0, x_29); x_41 = lean_array_mk(x_19); x_42 = lean_array_size(x_41); x_43 = 0; -x_44 = l_Array_mapMUnsafe_map___at_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__modelUsing____1___spec__3(x_42, x_43, x_41); +x_44 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_42, x_43, x_41); x_45 = lean_array_size(x_44); -x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_45, x_43, x_44); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_45, x_43, x_44); x_47 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_48 = l_Array_append___rarg(x_47, x_46); lean_dec(x_46); @@ -32256,7 +32278,7 @@ x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_29); lean_ctor_set(x_61, 1, x_60); x_62 = lean_mk_syntax_ident(x_18); -x_63 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1; +x_63 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1; lean_inc(x_29); lean_ctor_set_tag(x_30, 2); lean_ctor_set(x_30, 1, x_63); @@ -32264,9 +32286,9 @@ lean_ctor_set(x_30, 0, x_29); x_64 = lean_array_mk(x_19); x_65 = lean_array_size(x_64); x_66 = 0; -x_67 = l_Array_mapMUnsafe_map___at_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__modelUsing____1___spec__3(x_65, x_66, x_64); +x_67 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_65, x_66, x_64); x_68 = lean_array_size(x_67); -x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_68, x_66, x_67); +x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_68, x_66, x_67); x_70 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_71 = l_Array_append___rarg(x_70, x_69); lean_dec(x_69); @@ -32325,7 +32347,7 @@ if (lean_is_scalar(x_85)) { lean_ctor_set(x_87, 0, x_29); lean_ctor_set(x_87, 1, x_86); x_88 = lean_mk_syntax_ident(x_18); -x_89 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1; +x_89 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1; lean_inc(x_29); x_90 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_90, 0, x_29); @@ -32333,9 +32355,9 @@ lean_ctor_set(x_90, 1, x_89); x_91 = lean_array_mk(x_19); x_92 = lean_array_size(x_91); x_93 = 0; -x_94 = l_Array_mapMUnsafe_map___at_Std_DHashMap_Internal_Raw_u2080___aux__Std__Data__DHashMap__Internal__RawLemmas______macroRules__Std__DHashMap__Internal__Raw_u2080__tacticSimp__to__modelUsing____1___spec__3(x_92, x_93, x_91); +x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_92, x_93, x_91); x_95 = lean_array_size(x_94); -x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__1(x_95, x_93, x_94); +x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_95, x_93, x_94); x_97 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_98 = l_Array_append___rarg(x_97, x_96); lean_dec(x_96); @@ -32457,19 +32479,19 @@ lean_ctor_set_tag(x_208, 2); lean_ctor_set(x_208, 1, x_212); lean_ctor_set(x_208, 0, x_203); x_213 = lean_mk_syntax_ident(x_194); -x_214 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; +x_214 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; lean_inc(x_203); lean_ctor_set_tag(x_204, 2); lean_ctor_set(x_204, 1, x_214); lean_ctor_set(x_204, 0, x_203); x_215 = lean_mk_syntax_ident(x_197); -x_216 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; +x_216 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; lean_inc(x_203); lean_ctor_set_tag(x_198, 2); lean_ctor_set(x_198, 1, x_216); lean_ctor_set(x_198, 0, x_203); x_217 = lean_mk_syntax_ident(x_187); -x_218 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; +x_218 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; lean_inc(x_203); x_219 = l_Lean_Syntax_node3(x_203, x_218, x_215, x_198, x_217); x_220 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__3; @@ -32507,19 +32529,19 @@ x_232 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_232, 0, x_203); lean_ctor_set(x_232, 1, x_231); x_233 = lean_mk_syntax_ident(x_194); -x_234 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; +x_234 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; lean_inc(x_203); lean_ctor_set_tag(x_204, 2); lean_ctor_set(x_204, 1, x_234); lean_ctor_set(x_204, 0, x_203); x_235 = lean_mk_syntax_ident(x_197); -x_236 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; +x_236 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; lean_inc(x_203); lean_ctor_set_tag(x_198, 2); lean_ctor_set(x_198, 1, x_236); lean_ctor_set(x_198, 0, x_203); x_237 = lean_mk_syntax_ident(x_187); -x_238 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; +x_238 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; lean_inc(x_203); x_239 = l_Lean_Syntax_node3(x_203, x_238, x_235, x_198, x_237); x_240 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__3; @@ -32574,19 +32596,19 @@ if (lean_is_scalar(x_253)) { lean_ctor_set(x_255, 0, x_203); lean_ctor_set(x_255, 1, x_254); x_256 = lean_mk_syntax_ident(x_194); -x_257 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; +x_257 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; lean_inc(x_203); x_258 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_258, 0, x_203); lean_ctor_set(x_258, 1, x_257); x_259 = lean_mk_syntax_ident(x_197); -x_260 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; +x_260 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; lean_inc(x_203); lean_ctor_set_tag(x_198, 2); lean_ctor_set(x_198, 1, x_260); lean_ctor_set(x_198, 0, x_203); x_261 = lean_mk_syntax_ident(x_187); -x_262 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; +x_262 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; lean_inc(x_203); x_263 = l_Lean_Syntax_node3(x_203, x_262, x_259, x_198, x_261); x_264 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__3; @@ -32657,7 +32679,7 @@ if (lean_is_scalar(x_283)) { lean_ctor_set(x_285, 0, x_277); lean_ctor_set(x_285, 1, x_284); x_286 = lean_mk_syntax_ident(x_194); -x_287 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2; +x_287 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2; lean_inc(x_277); if (lean_is_scalar(x_280)) { x_288 = lean_alloc_ctor(2, 2, 0); @@ -32668,13 +32690,13 @@ if (lean_is_scalar(x_280)) { lean_ctor_set(x_288, 0, x_277); lean_ctor_set(x_288, 1, x_287); x_289 = lean_mk_syntax_ident(x_197); -x_290 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5; +x_290 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5; lean_inc(x_277); x_291 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_291, 0, x_277); lean_ctor_set(x_291, 1, x_290); x_292 = lean_mk_syntax_ident(x_187); -x_293 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4; +x_293 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4; lean_inc(x_277); x_294 = l_Lean_Syntax_node3(x_277, x_293, x_289, x_291, x_292); x_295 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__5___closed__3; @@ -32739,7 +32761,7 @@ lean_ctor_set(x_317, 1, x_321); lean_ctor_set(x_317, 0, x_312); x_322 = lean_array_size(x_185); x_323 = 0; -x_324 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_322, x_323, x_185); +x_324 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_322, x_323, x_185); x_325 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_326 = l_Array_append___rarg(x_325, x_324); lean_dec(x_324); @@ -32776,7 +32798,7 @@ lean_ctor_set(x_337, 0, x_312); lean_ctor_set(x_337, 1, x_336); x_338 = lean_array_size(x_185); x_339 = 0; -x_340 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_338, x_339, x_185); +x_340 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_338, x_339, x_185); x_341 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_342 = l_Array_append___rarg(x_341, x_340); lean_dec(x_340); @@ -32830,7 +32852,7 @@ lean_ctor_set(x_356, 0, x_312); lean_ctor_set(x_356, 1, x_355); x_357 = lean_array_size(x_185); x_358 = 0; -x_359 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_357, x_358, x_185); +x_359 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_357, x_358, x_185); x_360 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_361 = l_Array_append___rarg(x_360, x_359); lean_dec(x_359); @@ -32912,7 +32934,7 @@ x_27 = lean_box(0); x_28 = l_Lean_resolveNamespace___at_Lean_Elab_Command_elabExport___spec__1___closed__4; x_29 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__6___closed__1; lean_inc(x_1); -x_30 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_1, x_26, x_28, x_27, x_1, x_1, x_29, lean_box(0), x_2, x_3, x_4); +x_30 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4(x_1, x_26, x_28, x_27, x_1, x_1, x_29, lean_box(0), x_2, x_3, x_4); lean_dec(x_1); x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); @@ -33049,7 +33071,7 @@ lean_ctor_set(x_47, 1, x_51); lean_ctor_set(x_47, 0, x_42); x_52 = lean_array_size(x_34); x_53 = 0; -x_54 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_52, x_53, x_34); +x_54 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_52, x_53, x_34); x_55 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_56 = l_Array_append___rarg(x_55, x_54); lean_dec(x_54); @@ -33085,7 +33107,7 @@ lean_ctor_set(x_67, 0, x_42); lean_ctor_set(x_67, 1, x_66); x_68 = lean_array_size(x_34); x_69 = 0; -x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_68, x_69, x_34); +x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_68, x_69, x_34); x_71 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_72 = l_Array_append___rarg(x_71, x_70); lean_dec(x_70); @@ -33138,7 +33160,7 @@ lean_ctor_set(x_86, 0, x_42); lean_ctor_set(x_86, 1, x_85); x_87 = lean_array_size(x_34); x_88 = 0; -x_89 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__2(x_87, x_88, x_34); +x_89 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_87, x_88, x_34); x_90 = l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; x_91 = l_Array_append___rarg(x_90, x_89); lean_dec(x_89); @@ -33204,11 +33226,23 @@ x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls_ return x_6; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_5); @@ -36201,16 +36235,16 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1___clo if (builtin) {res = l___regBuiltin_Lean_Elab_Command_elabEoi_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__1); -l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__2); -l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__3); -l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__4); -l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__3___closed__5); +}l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__1); +l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__2); +l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__3); +l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__4); +l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5 = _init_l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Elab_Command_elabWhere_describeOpenDecls___spec__4___closed__5); l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__1 = _init_l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__1); l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__2 = _init_l_Lean_Elab_Command_elabWhere_describeOpenDecls___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/BuiltinNotation.c b/stage0/stdlib/Lean/Elab/BuiltinNotation.c index 39a3735c5193..e38d8e7a0a7a 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinNotation.c +++ b/stage0/stdlib/Lean/Elab/BuiltinNotation.c @@ -227,7 +227,7 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__42; static lean_object* l_Lean_Elab_Term_elabRunElab_unsafe__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange__1___closed__6; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetI__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetI___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*); @@ -856,7 +856,7 @@ static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elab uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__4___closed__1; static lean_object* l_Lean_Elab_Term_elabSubst___closed__7; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabRunElab__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandTypeAscription(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoeSortNotation__1___closed__2; @@ -3635,7 +3635,8 @@ x_30 = lean_ctor_get(x_27, 1); x_31 = lean_ctor_get(x_29, 0); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_environment_find(x_31, x_26); +x_32 = l_Lean_Environment_find_x3f(x_31, x_26); +lean_dec(x_26); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; @@ -3726,6 +3727,7 @@ lean_inc(x_54); lean_dec(x_52); lean_inc(x_48); x_55 = l_Lean_isPrivateNameFromImportedModule(x_54, x_48); +lean_dec(x_54); if (x_55 == 0) { lean_object* x_56; lean_object* x_57; @@ -3798,6 +3800,7 @@ lean_inc(x_70); lean_dec(x_68); lean_inc(x_48); x_71 = l_Lean_isPrivateNameFromImportedModule(x_70, x_48); +lean_dec(x_70); if (x_71 == 0) { lean_object* x_72; lean_object* x_73; @@ -3883,6 +3886,7 @@ lean_inc(x_90); lean_dec(x_87); lean_inc(x_85); x_91 = l_Lean_isPrivateNameFromImportedModule(x_90, x_85); +lean_dec(x_90); if (x_91 == 0) { lean_object* x_92; lean_object* x_93; @@ -4085,7 +4089,8 @@ lean_dec(x_27); x_135 = lean_ctor_get(x_133, 0); lean_inc(x_135); lean_dec(x_133); -x_136 = lean_environment_find(x_135, x_26); +x_136 = l_Lean_Environment_find_x3f(x_135, x_26); +lean_dec(x_26); if (lean_obj_tag(x_136) == 0) { lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; @@ -4184,6 +4189,7 @@ lean_inc(x_159); lean_dec(x_156); lean_inc(x_153); x_160 = l_Lean_isPrivateNameFromImportedModule(x_159, x_153); +lean_dec(x_159); if (x_160 == 0) { lean_object* x_161; lean_object* x_162; @@ -9353,7 +9359,8 @@ x_70 = lean_ctor_get(x_67, 1); x_71 = lean_ctor_get(x_69, 0); lean_inc(x_71); lean_dec(x_69); -x_72 = lean_environment_main_module(x_71); +x_72 = l_Lean_Environment_mainModule(x_71); +lean_dec(x_71); x_73 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; lean_inc(x_66); lean_inc(x_72); @@ -9461,7 +9468,8 @@ lean_dec(x_67); x_113 = lean_ctor_get(x_111, 0); lean_inc(x_113); lean_dec(x_111); -x_114 = lean_environment_main_module(x_113); +x_114 = l_Lean_Environment_mainModule(x_113); +lean_dec(x_113); x_115 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; lean_inc(x_66); lean_inc(x_114); @@ -9585,7 +9593,8 @@ x_35 = lean_ctor_get(x_33, 0); x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); lean_dec(x_35); -x_37 = lean_environment_main_module(x_36); +x_37 = l_Lean_Environment_mainModule(x_36); +lean_dec(x_36); x_38 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__7; x_39 = l_Lean_addMacroScope(x_37, x_38, x_32); x_40 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__6; @@ -9620,7 +9629,8 @@ lean_dec(x_33); x_49 = lean_ctor_get(x_47, 0); lean_inc(x_49); lean_dec(x_47); -x_50 = lean_environment_main_module(x_49); +x_50 = l_Lean_Environment_mainModule(x_49); +lean_dec(x_49); x_51 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__7; x_52 = l_Lean_addMacroScope(x_50, x_51, x_32); x_53 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__6; @@ -10376,7 +10386,8 @@ x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec(x_22); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__4; x_26 = l_Lean_addMacroScope(x_24, x_25, x_19); x_27 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__2; @@ -10406,7 +10417,8 @@ lean_dec(x_20); x_36 = lean_ctor_get(x_34, 0); lean_inc(x_36); lean_dec(x_34); -x_37 = lean_environment_main_module(x_36); +x_37 = l_Lean_Environment_mainModule(x_36); +lean_dec(x_36); x_38 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__4; x_39 = l_Lean_addMacroScope(x_37, x_38, x_19); x_40 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__2; @@ -10519,7 +10531,8 @@ if (lean_is_exclusive(x_72)) { x_76 = lean_ctor_get(x_73, 0); lean_inc(x_76); lean_dec(x_73); -x_77 = lean_environment_main_module(x_76); +x_77 = l_Lean_Environment_mainModule(x_76); +lean_dec(x_76); x_78 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__4; x_79 = l_Lean_addMacroScope(x_77, x_78, x_71); x_80 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__2; @@ -11238,7 +11251,8 @@ lean_dec(x_29); x_32 = lean_ctor_get(x_30, 0); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_environment_main_module(x_32); +x_33 = l_Lean_Environment_mainModule(x_32); +lean_dec(x_32); x_34 = l_Lean_Elab_Term_elabPanic___closed__5; x_35 = l_Lean_addMacroScope(x_33, x_34, x_28); x_36 = l_Lean_Elab_Term_elabPanic___closed__4; @@ -11249,7 +11263,8 @@ lean_ctor_set(x_38, 0, x_27); lean_ctor_set(x_38, 1, x_36); lean_ctor_set(x_38, 2, x_35); lean_ctor_set(x_38, 3, x_37); -x_39 = lean_environment_main_module(x_21); +x_39 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_40 = 1; x_41 = l_Lean_Elab_Term_elabPanic___closed__8; x_42 = l_Lean_Name_toString(x_39, x_40, x_41); @@ -11298,7 +11313,8 @@ lean_dec(x_62); x_65 = lean_ctor_get(x_63, 0); lean_inc(x_65); lean_dec(x_63); -x_66 = lean_environment_main_module(x_65); +x_66 = l_Lean_Environment_mainModule(x_65); +lean_dec(x_65); x_67 = l_Lean_Elab_Term_elabPanic___closed__11; x_68 = l_Lean_addMacroScope(x_66, x_67, x_61); x_69 = l_Lean_Elab_Term_elabPanic___closed__10; @@ -11309,7 +11325,8 @@ lean_ctor_set(x_71, 0, x_60); lean_ctor_set(x_71, 1, x_69); lean_ctor_set(x_71, 2, x_68); lean_ctor_set(x_71, 3, x_70); -x_72 = lean_environment_main_module(x_21); +x_72 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_73 = 1; x_74 = l_Lean_Elab_Term_elabPanic___closed__8; x_75 = l_Lean_Name_toString(x_72, x_73, x_74); @@ -16028,7 +16045,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -16210,7 +16228,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); diff --git a/stage0/stdlib/Lean/Elab/BuiltinTerm.c b/stage0/stdlib/Lean/Elab/BuiltinTerm.c index 68580fc6f7be..d5738f132ded 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinTerm.c +++ b/stage0/stdlib/Lean/Elab/BuiltinTerm.c @@ -232,7 +232,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Term_e static lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3___closed__1; lean_object* l_Lean_ScopedEnvExtension_pushScope___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_declRange__1___closed__6; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_elabClear___spec__19(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_elabOpen___spec__41(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabEnsureExpectedType_declRange__1___closed__6; @@ -315,14 +315,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinTerm_0__Lean_Elab_Term_mkF LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_elabClear___spec__22(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeStx__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabHole_declRange__1___closed__5; -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabCompletion__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabScientificLit__1___closed__1; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeContainsMVar_declRange__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPatternErr__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSort_declRange__1___closed__2; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabRawNatLit__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabRawNatLit__1___closed__3; uint8_t l_Lean_Expr_hasMVar(lean_object*); @@ -451,6 +449,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabWithDeclName__1___closed__ lean_object* l_Lean_Elab_Term_isTacticOrPostponedHole_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_declRange__1___closed__5; +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); static lean_object* l_Lean_pushScope___at_Lean_Elab_Term_elabOpen___spec__1___closed__1; @@ -984,6 +983,7 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabDeclName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabCharLit__1___closed__2; lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSetOption_declRange__1___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_elabClear___spec__25(lean_object*, lean_object*, size_t, size_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabQuotedName_declRange__1___closed__3; @@ -18277,8 +18277,7 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_1); -x_15 = lean_environment_find(x_14, x_1); +x_15 = l_Lean_Environment_find_x3f(x_14, x_1); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -18318,8 +18317,7 @@ lean_dec(x_10); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_1); -x_27 = lean_environment_find(x_26, x_1); +x_27 = l_Lean_Environment_find_x3f(x_26, x_1); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; @@ -25276,7 +25274,7 @@ lean_object* x_28; uint8_t x_29; uint8_t x_30; x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -x_29 = l_Lean_Kernel_isDiagnosticsEnabled(x_28); +x_29 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_28); lean_dec(x_28); if (x_29 == 0) { @@ -25330,7 +25328,7 @@ lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean x_35 = lean_ctor_get(x_32, 0); x_36 = lean_ctor_get(x_32, 4); lean_dec(x_36); -x_37 = l_Lean_Kernel_enableDiag(x_35, x_23); +x_37 = l_Lean_Kernel_Environment_enableDiag(x_35, x_23); x_38 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabOpen___spec__2___closed__3; lean_ctor_set(x_32, 4, x_38); lean_ctor_set(x_32, 0, x_37); @@ -25360,7 +25358,7 @@ lean_inc(x_45); lean_inc(x_44); lean_inc(x_43); lean_dec(x_32); -x_50 = l_Lean_Kernel_enableDiag(x_43, x_23); +x_50 = l_Lean_Kernel_Environment_enableDiag(x_43, x_23); x_51 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabOpen___spec__2___closed__3; x_52 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_52, 0, x_50); diff --git a/stage0/stdlib/Lean/Elab/Calc.c b/stage0/stdlib/Lean/Elab/Calc.c index 075301bc5e1c..12608bb25c28 100644 --- a/stage0/stdlib/Lean/Elab/Calc.c +++ b/stage0/stdlib/Lean/Elab/Calc.c @@ -192,7 +192,7 @@ static lean_object* l_Lean_Elab_Term_annotateFirstHoleWithType_go___lambda__1___ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_annotateFirstHoleWithType_go___spec__14(lean_object*, size_t, size_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getCalcRelation_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_mkCalcFirstStepView___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_annotateFirstHoleWithType_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -16049,7 +16049,8 @@ x_41 = lean_ctor_get(x_39, 0); x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); lean_dec(x_41); -x_43 = lean_environment_main_module(x_42); +x_43 = l_Lean_Environment_mainModule(x_42); +lean_dec(x_42); x_44 = l_Lean_Elab_Term_mkCalcFirstStepView___closed__10; x_45 = l_Lean_addMacroScope(x_43, x_44, x_13); x_46 = l_Lean_Elab_Term_mkCalcFirstStepView___closed__9; @@ -16077,7 +16078,8 @@ lean_dec(x_39); x_52 = lean_ctor_get(x_50, 0); lean_inc(x_52); lean_dec(x_50); -x_53 = lean_environment_main_module(x_52); +x_53 = l_Lean_Environment_mainModule(x_52); +lean_dec(x_52); x_54 = l_Lean_Elab_Term_mkCalcFirstStepView___closed__10; x_55 = l_Lean_addMacroScope(x_53, x_54, x_13); x_56 = l_Lean_Elab_Term_mkCalcFirstStepView___closed__9; @@ -16135,7 +16137,8 @@ if (lean_is_exclusive(x_70)) { x_74 = lean_ctor_get(x_71, 0); lean_inc(x_74); lean_dec(x_71); -x_75 = lean_environment_main_module(x_74); +x_75 = l_Lean_Environment_mainModule(x_74); +lean_dec(x_74); x_76 = l_Lean_Elab_Term_mkCalcFirstStepView___closed__10; x_77 = l_Lean_addMacroScope(x_75, x_76, x_13); x_78 = l_Lean_Elab_Term_mkCalcFirstStepView___closed__9; diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index b2e4c9c41174..b723211c6601 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -46,7 +46,6 @@ lean_object* l_EIO_toBaseIO___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__5; static lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Elab_Command_wrapAsyncAsSnapshot___spec__10___closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_wrapAsyncAsSnapshot___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_diagExt; lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); uint8_t l_Lean_Elab_isAbortExceptionId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__5___boxed(lean_object*); @@ -220,7 +219,6 @@ static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed_ lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__2___closed__2; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; -lean_object* l_Lean_EnvExtension_modifyState___rarg(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___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_wrapAsyncAsSnapshot___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_format(lean_object*, lean_object*, lean_object*); @@ -299,7 +297,6 @@ LEAN_EXPORT lean_object* l_Array_zipWithAux___at_Lean_Elab_Command_elabCommand__ LEAN_EXPORT lean_object* l_Lean_liftCommandElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__2___lambda__3(lean_object*, 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__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___closed__1; @@ -313,7 +310,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_runLinte LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__2___rarg___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___auto____x40_Lean_Elab_Command___hyg_472____closed__21; LEAN_EXPORT lean_object* l_Lean_Elab_withLogging___at_Lean_Elab_Command_withLoggingExceptions___spec__1(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_addTraceAsMessages___at_Lean_Elab_Command_wrapAsyncAsSnapshot___spec__3___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM; @@ -324,7 +320,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLintersAsync___lambda__2___boxed 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*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__4; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_wrapAsyncAsSnapshot___spec__9___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withoutCommandIncrementality___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -444,6 +439,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___h static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_resetDiag(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_liftCommandElabMCore(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_withSetOptionIn___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -463,6 +459,7 @@ static lean_object* l_Lean_withSetOptionIn___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___closed__3; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__1; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__2; @@ -481,7 +478,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___boxed(lean_object*, lean_object*, 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*); extern lean_object* l_Lean_Elab_async; -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*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withScope(lean_object*); @@ -745,7 +741,7 @@ LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_runLinters___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_addTraceAsMessages___spec__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Command_withMacroExpansion___spec__1(lean_object*); 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*); @@ -850,7 +846,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCom LEAN_EXPORT lean_object* l_Lean_Elab_Command_wrapAsyncAsSnapshot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabCommand___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_iotaTR(lean_object*); -lean_object* l_Lean_Kernel_resetDiag___lambda__1(lean_object*); 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*); @@ -955,6 +950,7 @@ static lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___closed lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l_Lean_Elab_Command_wrapAsyncAsSnapshot___lambda__5___closed__9; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_647____closed__13; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_showPartialSyntaxErrors; @@ -3808,22 +3804,6 @@ return x_28; static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1() { _start: { -lean_object* x_1; -x_1 = l_Lean_diagExt; -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Kernel_resetDiag___lambda__1), 1, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3() { -_start: -{ lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Elab_Command_mkState___closed__8; x_2 = lean_alloc_ctor(0, 2, 0); @@ -3832,7 +3812,7 @@ lean_ctor_set(x_2, 1, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4() { +static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -3845,7 +3825,7 @@ lean_ctor_set(x_3, 2, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__5() { +static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3() { _start: { lean_object* x_1; @@ -3856,7 +3836,7 @@ return x_1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; 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; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; uint8_t x_37; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; uint8_t x_35; x_5 = lean_st_ref_get(x_3, x_4); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); @@ -3886,661 +3866,659 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_6, 8); lean_inc(x_18); lean_dec(x_6); -x_19 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1; -x_20 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; -x_21 = l_Lean_EnvExtension_modifyState___rarg(x_19, x_11, x_20); -x_22 = l_Lean_Elab_Command_instInhabitedScope; -x_23 = l_List_head_x21___rarg(x_22, x_12); +x_19 = l_Lean_Kernel_Environment_resetDiag(x_11); +x_20 = l_Lean_Elab_Command_instInhabitedScope; +x_21 = l_List_head_x21___rarg(x_20, x_12); lean_dec(x_12); -x_24 = lean_ctor_get(x_2, 0); -x_25 = lean_ctor_get(x_2, 1); -x_26 = lean_ctor_get(x_2, 2); -x_27 = lean_ctor_get(x_2, 5); -x_28 = lean_ctor_get(x_2, 6); -x_29 = lean_ctor_get(x_2, 9); -x_30 = lean_ctor_get_uint8(x_2, sizeof(void*)*10); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -x_32 = lean_ctor_get(x_23, 2); -lean_inc(x_32); -x_33 = lean_ctor_get(x_23, 3); -lean_inc(x_33); -lean_dec(x_23); -x_34 = l_Lean_Core_getMaxHeartbeats(x_31); -x_35 = 0; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +x_25 = lean_ctor_get(x_2, 5); +x_26 = lean_ctor_get(x_2, 6); +x_27 = lean_ctor_get(x_2, 9); +x_28 = lean_ctor_get_uint8(x_2, sizeof(void*)*10); +x_29 = lean_ctor_get(x_21, 1); lean_inc(x_29); -lean_inc(x_27); -lean_inc(x_28); -lean_inc(x_26); +x_30 = lean_ctor_get(x_21, 2); +lean_inc(x_30); +x_31 = lean_ctor_get(x_21, 3); lean_inc(x_31); +lean_dec(x_21); +x_32 = l_Lean_Core_getMaxHeartbeats(x_29); +x_33 = 0; +lean_inc(x_27); lean_inc(x_25); +lean_inc(x_26); lean_inc(x_24); -x_36 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_36, 0, x_24); -lean_ctor_set(x_36, 1, x_25); -lean_ctor_set(x_36, 2, x_31); -lean_ctor_set(x_36, 3, x_26); -lean_ctor_set(x_36, 4, x_14); -lean_ctor_set(x_36, 5, x_28); -lean_ctor_set(x_36, 6, x_32); -lean_ctor_set(x_36, 7, x_33); -lean_ctor_set(x_36, 8, x_9); -lean_ctor_set(x_36, 9, x_34); -lean_ctor_set(x_36, 10, x_27); -lean_ctor_set(x_36, 11, x_29); -lean_ctor_set_uint8(x_36, sizeof(void*)*12, x_35); -lean_ctor_set_uint8(x_36, sizeof(void*)*12 + 1, x_30); -x_37 = !lean_is_exclusive(x_16); -if (x_37 == 0) +lean_inc(x_29); +lean_inc(x_23); +lean_inc(x_22); +x_34 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_23); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_24); +lean_ctor_set(x_34, 4, x_14); +lean_ctor_set(x_34, 5, x_26); +lean_ctor_set(x_34, 6, x_30); +lean_ctor_set(x_34, 7, x_31); +lean_ctor_set(x_34, 8, x_9); +lean_ctor_set(x_34, 9, x_32); +lean_ctor_set(x_34, 10, x_25); +lean_ctor_set(x_34, 11, x_27); +lean_ctor_set_uint8(x_34, sizeof(void*)*12, x_33); +lean_ctor_set_uint8(x_34, sizeof(void*)*12 + 1, x_28); +x_35 = !lean_is_exclusive(x_16); +if (x_35 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; -x_38 = lean_ctor_get(x_16, 1); -lean_dec(x_38); -x_39 = lean_ctor_get(x_16, 0); -lean_dec(x_39); -x_40 = l_Lean_Elab_Command_mkState___closed__8; -x_41 = l_Lean_Elab_Command_mkState___closed__11; -lean_ctor_set(x_16, 1, x_41); -lean_ctor_set(x_16, 0, x_40); -x_42 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3; -x_43 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; -x_44 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_44, 0, x_21); -lean_ctor_set(x_44, 1, x_13); -lean_ctor_set(x_44, 2, x_15); -lean_ctor_set(x_44, 3, x_17); -lean_ctor_set(x_44, 4, x_42); -lean_ctor_set(x_44, 5, x_43); -lean_ctor_set(x_44, 6, x_16); -lean_ctor_set(x_44, 7, x_18); -x_128 = lean_st_mk_ref(x_44, x_10); -x_129 = lean_ctor_get(x_128, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_128, 1); -lean_inc(x_130); -lean_dec(x_128); -x_131 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__5; -x_132 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_31, x_131); -x_133 = lean_st_ref_get(x_129, x_130); -x_134 = lean_ctor_get(x_133, 0); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_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; uint8_t x_135; uint8_t x_136; +x_36 = lean_ctor_get(x_16, 1); +lean_dec(x_36); +x_37 = lean_ctor_get(x_16, 0); +lean_dec(x_37); +x_38 = l_Lean_Elab_Command_mkState___closed__8; +x_39 = l_Lean_Elab_Command_mkState___closed__11; +lean_ctor_set(x_16, 1, x_39); +lean_ctor_set(x_16, 0, x_38); +x_40 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1; +x_41 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; +x_42 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_42, 0, x_19); +lean_ctor_set(x_42, 1, x_13); +lean_ctor_set(x_42, 2, x_15); +lean_ctor_set(x_42, 3, x_17); +lean_ctor_set(x_42, 4, x_40); +lean_ctor_set(x_42, 5, x_41); +lean_ctor_set(x_42, 6, x_16); +lean_ctor_set(x_42, 7, x_18); +x_126 = lean_st_mk_ref(x_42, x_10); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3; +x_130 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_29, x_129); +x_131 = lean_st_ref_get(x_127, x_128); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_131, 1); +lean_inc(x_133); +lean_dec(x_131); +x_134 = lean_ctor_get(x_132, 0); lean_inc(x_134); -x_135 = lean_ctor_get(x_133, 1); -lean_inc(x_135); -lean_dec(x_133); -x_136 = lean_ctor_get(x_134, 0); -lean_inc(x_136); +lean_dec(x_132); +x_135 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_134); lean_dec(x_134); -x_137 = l_Lean_Kernel_isDiagnosticsEnabled(x_136); -lean_dec(x_136); -if (x_137 == 0) +if (x_135 == 0) { -if (x_132 == 0) +if (x_130 == 0) { -uint8_t x_203; -x_203 = 1; -x_138 = x_203; -goto block_202; +uint8_t x_201; +x_201 = 1; +x_136 = x_201; +goto block_200; } else { -x_138 = x_35; -goto block_202; +x_136 = x_33; +goto block_200; } } else { -if (x_132 == 0) +if (x_130 == 0) { -x_138 = x_35; -goto block_202; +x_136 = x_33; +goto block_200; } else { -uint8_t x_204; -x_204 = 1; -x_138 = x_204; -goto block_202; +uint8_t x_202; +x_202 = 1; +x_136 = x_202; +goto block_200; } } -block_127: +block_125: { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_47 = lean_ctor_get(x_45, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_45, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_dec(x_43); +x_47 = lean_st_ref_take(x_3, x_44); +x_48 = lean_ctor_get(x_47, 0); lean_inc(x_48); -lean_dec(x_45); -x_49 = lean_st_ref_take(x_3, x_46); -x_50 = lean_ctor_get(x_49, 0); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = lean_ctor_get(x_46, 0); lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); +x_51 = lean_ctor_get(x_46, 1); lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_ctor_get(x_48, 0); +x_52 = lean_ctor_get(x_46, 2); lean_inc(x_52); -x_53 = lean_ctor_get(x_48, 1); +x_53 = lean_ctor_get(x_46, 3); lean_inc(x_53); -x_54 = lean_ctor_get(x_48, 2); +x_54 = lean_ctor_get(x_46, 5); lean_inc(x_54); -x_55 = lean_ctor_get(x_48, 3); +x_55 = lean_ctor_get(x_46, 6); lean_inc(x_55); -x_56 = lean_ctor_get(x_48, 5); +x_56 = lean_ctor_get(x_46, 7); lean_inc(x_56); -x_57 = lean_ctor_get(x_48, 6); -lean_inc(x_57); -x_58 = lean_ctor_get(x_48, 7); -lean_inc(x_58); -lean_dec(x_48); -x_59 = !lean_is_exclusive(x_50); -if (x_59 == 0) +lean_dec(x_46); +x_57 = !lean_is_exclusive(x_48); +if (x_57 == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_60 = lean_ctor_get(x_50, 1); -x_61 = lean_ctor_get(x_50, 6); -x_62 = lean_ctor_get(x_50, 7); -x_63 = lean_ctor_get(x_50, 8); +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_58 = lean_ctor_get(x_48, 1); +x_59 = lean_ctor_get(x_48, 6); +x_60 = lean_ctor_get(x_48, 7); +x_61 = lean_ctor_get(x_48, 8); +lean_dec(x_61); +x_62 = lean_ctor_get(x_48, 5); +lean_dec(x_62); +x_63 = lean_ctor_get(x_48, 3); lean_dec(x_63); -x_64 = lean_ctor_get(x_50, 5); +x_64 = lean_ctor_get(x_48, 0); lean_dec(x_64); -x_65 = lean_ctor_get(x_50, 3); -lean_dec(x_65); -x_66 = lean_ctor_get(x_50, 0); -lean_dec(x_66); -x_67 = l_Lean_MessageLog_append(x_60, x_56); -x_68 = !lean_is_exclusive(x_61); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_69 = lean_ctor_get(x_61, 1); -x_70 = lean_ctor_get(x_57, 1); -lean_inc(x_70); -lean_dec(x_57); -x_71 = l_Lean_PersistentArray_append___rarg(x_69, x_70); -lean_dec(x_70); -lean_ctor_set(x_61, 1, x_71); -x_72 = !lean_is_exclusive(x_62); -if (x_72 == 0) +x_65 = l_Lean_MessageLog_append(x_58, x_54); +x_66 = !lean_is_exclusive(x_59); +if (x_66 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_73 = lean_ctor_get(x_62, 0); -lean_dec(x_73); -x_74 = lean_ctor_get(x_55, 0); -lean_inc(x_74); +lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +x_67 = lean_ctor_get(x_59, 1); +x_68 = lean_ctor_get(x_55, 1); +lean_inc(x_68); lean_dec(x_55); -x_75 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_28, x_74); -lean_ctor_set(x_62, 0, x_75); -lean_ctor_set(x_50, 8, x_58); -lean_ctor_set(x_50, 5, x_54); -lean_ctor_set(x_50, 3, x_53); -lean_ctor_set(x_50, 1, x_67); -lean_ctor_set(x_50, 0, x_52); -x_76 = lean_st_ref_set(x_3, x_50, x_51); -x_77 = !lean_is_exclusive(x_76); -if (x_77 == 0) +x_69 = l_Lean_PersistentArray_append___rarg(x_67, x_68); +lean_dec(x_68); +lean_ctor_set(x_59, 1, x_69); +x_70 = !lean_is_exclusive(x_60); +if (x_70 == 0) { -lean_object* x_78; -x_78 = lean_ctor_get(x_76, 0); -lean_dec(x_78); -lean_ctor_set(x_76, 0, x_47); -return x_76; +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_71 = lean_ctor_get(x_60, 0); +lean_dec(x_71); +x_72 = lean_ctor_get(x_53, 0); +lean_inc(x_72); +lean_dec(x_53); +x_73 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_26, x_72); +lean_ctor_set(x_60, 0, x_73); +lean_ctor_set(x_48, 8, x_56); +lean_ctor_set(x_48, 5, x_52); +lean_ctor_set(x_48, 3, x_51); +lean_ctor_set(x_48, 1, x_65); +lean_ctor_set(x_48, 0, x_50); +x_74 = lean_st_ref_set(x_3, x_48, x_49); +x_75 = !lean_is_exclusive(x_74); +if (x_75 == 0) +{ +lean_object* x_76; +x_76 = lean_ctor_get(x_74, 0); +lean_dec(x_76); +lean_ctor_set(x_74, 0, x_45); +return x_74; } else { -lean_object* x_79; lean_object* x_80; -x_79 = lean_ctor_get(x_76, 1); -lean_inc(x_79); -lean_dec(x_76); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_47); -lean_ctor_set(x_80, 1, x_79); -return x_80; +lean_object* x_77; lean_object* x_78; +x_77 = lean_ctor_get(x_74, 1); +lean_inc(x_77); +lean_dec(x_74); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_45); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } else { -uint64_t x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_81 = lean_ctor_get_uint64(x_62, sizeof(void*)*1); -lean_dec(x_62); -x_82 = lean_ctor_get(x_55, 0); -lean_inc(x_82); -lean_dec(x_55); -x_83 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_28, x_82); -x_84 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set_uint64(x_84, sizeof(void*)*1, x_81); -lean_ctor_set(x_50, 8, x_58); -lean_ctor_set(x_50, 7, x_84); -lean_ctor_set(x_50, 5, x_54); -lean_ctor_set(x_50, 3, x_53); -lean_ctor_set(x_50, 1, x_67); -lean_ctor_set(x_50, 0, x_52); -x_85 = lean_st_ref_set(x_3, x_50, x_51); -x_86 = lean_ctor_get(x_85, 1); -lean_inc(x_86); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_87 = x_85; +uint64_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_79 = lean_ctor_get_uint64(x_60, sizeof(void*)*1); +lean_dec(x_60); +x_80 = lean_ctor_get(x_53, 0); +lean_inc(x_80); +lean_dec(x_53); +x_81 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_26, x_80); +x_82 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set_uint64(x_82, sizeof(void*)*1, x_79); +lean_ctor_set(x_48, 8, x_56); +lean_ctor_set(x_48, 7, x_82); +lean_ctor_set(x_48, 5, x_52); +lean_ctor_set(x_48, 3, x_51); +lean_ctor_set(x_48, 1, x_65); +lean_ctor_set(x_48, 0, x_50); +x_83 = lean_st_ref_set(x_3, x_48, x_49); +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_85 = x_83; } else { - lean_dec_ref(x_85); - x_87 = lean_box(0); + lean_dec_ref(x_83); + x_85 = lean_box(0); } -if (lean_is_scalar(x_87)) { - x_88 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(0, 2, 0); } else { - x_88 = x_87; + x_86 = x_85; } -lean_ctor_set(x_88, 0, x_47); -lean_ctor_set(x_88, 1, x_86); -return x_88; +lean_ctor_set(x_86, 0, x_45); +lean_ctor_set(x_86, 1, x_84); +return x_86; } } else { -uint8_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint64_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_89 = lean_ctor_get_uint8(x_61, sizeof(void*)*2); -x_90 = lean_ctor_get(x_61, 0); -x_91 = lean_ctor_get(x_61, 1); -lean_inc(x_91); +uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint64_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_87 = lean_ctor_get_uint8(x_59, sizeof(void*)*2); +x_88 = lean_ctor_get(x_59, 0); +x_89 = lean_ctor_get(x_59, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_59); +x_90 = lean_ctor_get(x_55, 1); lean_inc(x_90); -lean_dec(x_61); -x_92 = lean_ctor_get(x_57, 1); -lean_inc(x_92); -lean_dec(x_57); -x_93 = l_Lean_PersistentArray_append___rarg(x_91, x_92); -lean_dec(x_92); -x_94 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_94, 0, x_90); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_uint8(x_94, sizeof(void*)*2, x_89); -x_95 = lean_ctor_get_uint64(x_62, sizeof(void*)*1); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - x_96 = x_62; +lean_dec(x_55); +x_91 = l_Lean_PersistentArray_append___rarg(x_89, x_90); +lean_dec(x_90); +x_92 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_92, 0, x_88); +lean_ctor_set(x_92, 1, x_91); +lean_ctor_set_uint8(x_92, sizeof(void*)*2, x_87); +x_93 = lean_ctor_get_uint64(x_60, sizeof(void*)*1); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + x_94 = x_60; } else { - lean_dec_ref(x_62); - x_96 = lean_box(0); + lean_dec_ref(x_60); + x_94 = lean_box(0); } -x_97 = lean_ctor_get(x_55, 0); -lean_inc(x_97); -lean_dec(x_55); -x_98 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_28, x_97); -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_96; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_95); -lean_ctor_set(x_50, 8, x_58); -lean_ctor_set(x_50, 7, x_99); -lean_ctor_set(x_50, 6, x_94); -lean_ctor_set(x_50, 5, x_54); -lean_ctor_set(x_50, 3, x_53); -lean_ctor_set(x_50, 1, x_67); -lean_ctor_set(x_50, 0, x_52); -x_100 = lean_st_ref_set(x_3, x_50, x_51); -x_101 = lean_ctor_get(x_100, 1); -lean_inc(x_101); -if (lean_is_exclusive(x_100)) { - lean_ctor_release(x_100, 0); - lean_ctor_release(x_100, 1); - x_102 = x_100; +x_95 = lean_ctor_get(x_53, 0); +lean_inc(x_95); +lean_dec(x_53); +x_96 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_26, x_95); +if (lean_is_scalar(x_94)) { + x_97 = lean_alloc_ctor(0, 1, 8); } else { - lean_dec_ref(x_100); - x_102 = lean_box(0); + x_97 = x_94; } -if (lean_is_scalar(x_102)) { - x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set_uint64(x_97, sizeof(void*)*1, x_93); +lean_ctor_set(x_48, 8, x_56); +lean_ctor_set(x_48, 7, x_97); +lean_ctor_set(x_48, 6, x_92); +lean_ctor_set(x_48, 5, x_52); +lean_ctor_set(x_48, 3, x_51); +lean_ctor_set(x_48, 1, x_65); +lean_ctor_set(x_48, 0, x_50); +x_98 = lean_st_ref_set(x_3, x_48, x_49); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; } else { - x_103 = x_102; + lean_dec_ref(x_98); + x_100 = lean_box(0); } -lean_ctor_set(x_103, 0, x_47); -lean_ctor_set(x_103, 1, x_101); -return x_103; +if (lean_is_scalar(x_100)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_100; +} +lean_ctor_set(x_101, 0, x_45); +lean_ctor_set(x_101, 1, x_99); +return x_101; } } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; uint64_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_104 = lean_ctor_get(x_50, 1); -x_105 = lean_ctor_get(x_50, 2); -x_106 = lean_ctor_get(x_50, 4); -x_107 = lean_ctor_get(x_50, 6); -x_108 = lean_ctor_get(x_50, 7); -lean_inc(x_108); -lean_inc(x_107); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint64_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_102 = lean_ctor_get(x_48, 1); +x_103 = lean_ctor_get(x_48, 2); +x_104 = lean_ctor_get(x_48, 4); +x_105 = lean_ctor_get(x_48, 6); +x_106 = lean_ctor_get(x_48, 7); lean_inc(x_106); lean_inc(x_105); lean_inc(x_104); -lean_dec(x_50); -x_109 = l_Lean_MessageLog_append(x_104, x_56); -x_110 = lean_ctor_get_uint8(x_107, sizeof(void*)*2); -x_111 = lean_ctor_get(x_107, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_107, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_48); +x_107 = l_Lean_MessageLog_append(x_102, x_54); +x_108 = lean_ctor_get_uint8(x_105, sizeof(void*)*2); +x_109 = lean_ctor_get(x_105, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_105, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_111 = x_105; +} else { + lean_dec_ref(x_105); + x_111 = lean_box(0); +} +x_112 = lean_ctor_get(x_55, 1); lean_inc(x_112); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_113 = x_107; +lean_dec(x_55); +x_113 = l_Lean_PersistentArray_append___rarg(x_110, x_112); +lean_dec(x_112); +if (lean_is_scalar(x_111)) { + x_114 = lean_alloc_ctor(0, 2, 1); } else { - lean_dec_ref(x_107); - x_113 = lean_box(0); + x_114 = x_111; } -x_114 = lean_ctor_get(x_57, 1); -lean_inc(x_114); -lean_dec(x_57); -x_115 = l_Lean_PersistentArray_append___rarg(x_112, x_114); -lean_dec(x_114); -if (lean_is_scalar(x_113)) { - x_116 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_114, 0, x_109); +lean_ctor_set(x_114, 1, x_113); +lean_ctor_set_uint8(x_114, sizeof(void*)*2, x_108); +x_115 = lean_ctor_get_uint64(x_106, sizeof(void*)*1); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + x_116 = x_106; } else { - x_116 = x_113; + lean_dec_ref(x_106); + x_116 = lean_box(0); } -lean_ctor_set(x_116, 0, x_111); -lean_ctor_set(x_116, 1, x_115); -lean_ctor_set_uint8(x_116, sizeof(void*)*2, x_110); -x_117 = lean_ctor_get_uint64(x_108, sizeof(void*)*1); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - x_118 = x_108; +x_117 = lean_ctor_get(x_53, 0); +lean_inc(x_117); +lean_dec(x_53); +x_118 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_26, x_117); +if (lean_is_scalar(x_116)) { + x_119 = lean_alloc_ctor(0, 1, 8); } else { - lean_dec_ref(x_108); - x_118 = lean_box(0); + x_119 = x_116; } -x_119 = lean_ctor_get(x_55, 0); -lean_inc(x_119); -lean_dec(x_55); -x_120 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_28, x_119); -if (lean_is_scalar(x_118)) { - x_121 = lean_alloc_ctor(0, 1, 8); -} else { - x_121 = x_118; -} -lean_ctor_set(x_121, 0, x_120); -lean_ctor_set_uint64(x_121, sizeof(void*)*1, x_117); -x_122 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_122, 0, x_52); -lean_ctor_set(x_122, 1, x_109); -lean_ctor_set(x_122, 2, x_105); -lean_ctor_set(x_122, 3, x_53); -lean_ctor_set(x_122, 4, x_106); -lean_ctor_set(x_122, 5, x_54); -lean_ctor_set(x_122, 6, x_116); -lean_ctor_set(x_122, 7, x_121); -lean_ctor_set(x_122, 8, x_58); -x_123 = lean_st_ref_set(x_3, x_122, x_51); -x_124 = lean_ctor_get(x_123, 1); -lean_inc(x_124); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_125 = x_123; +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set_uint64(x_119, sizeof(void*)*1, x_115); +x_120 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_120, 0, x_50); +lean_ctor_set(x_120, 1, x_107); +lean_ctor_set(x_120, 2, x_103); +lean_ctor_set(x_120, 3, x_51); +lean_ctor_set(x_120, 4, x_104); +lean_ctor_set(x_120, 5, x_52); +lean_ctor_set(x_120, 6, x_114); +lean_ctor_set(x_120, 7, x_119); +lean_ctor_set(x_120, 8, x_56); +x_121 = lean_st_ref_set(x_3, x_120, x_49); +x_122 = lean_ctor_get(x_121, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_121)) { + lean_ctor_release(x_121, 0); + lean_ctor_release(x_121, 1); + x_123 = x_121; } else { - lean_dec_ref(x_123); - x_125 = lean_box(0); + lean_dec_ref(x_121); + x_123 = lean_box(0); } -if (lean_is_scalar(x_125)) { - x_126 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_123)) { + x_124 = lean_alloc_ctor(0, 2, 0); } else { - x_126 = x_125; + x_124 = x_123; } -lean_ctor_set(x_126, 0, x_47); -lean_ctor_set(x_126, 1, x_124); -return x_126; +lean_ctor_set(x_124, 0, x_45); +lean_ctor_set(x_124, 1, x_122); +return x_124; } } -block_202: +block_200: { -if (x_138 == 0) +if (x_136 == 0) { -lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; -x_139 = lean_st_ref_take(x_129, x_135); -x_140 = lean_ctor_get(x_139, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_139, 1); -lean_inc(x_141); -lean_dec(x_139); -x_142 = !lean_is_exclusive(x_140); -if (x_142 == 0) +lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; +x_137 = lean_st_ref_take(x_127, x_133); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_137, 1); +lean_inc(x_139); +lean_dec(x_137); +x_140 = !lean_is_exclusive(x_138); +if (x_140 == 0) { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_143 = lean_ctor_get(x_140, 0); -x_144 = lean_ctor_get(x_140, 4); +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_141 = lean_ctor_get(x_138, 0); +x_142 = lean_ctor_get(x_138, 4); +lean_dec(x_142); +x_143 = l_Lean_Kernel_Environment_enableDiag(x_141, x_130); +lean_ctor_set(x_138, 4, x_40); +lean_ctor_set(x_138, 0, x_143); +x_144 = lean_st_ref_set(x_127, x_138, x_139); +x_145 = lean_ctor_get(x_144, 1); +lean_inc(x_145); lean_dec(x_144); -x_145 = l_Lean_Kernel_enableDiag(x_143, x_132); -lean_ctor_set(x_140, 4, x_42); -lean_ctor_set(x_140, 0, x_145); -x_146 = lean_st_ref_set(x_129, x_140, x_141); -x_147 = lean_ctor_get(x_146, 1); -lean_inc(x_147); -lean_dec(x_146); -x_148 = lean_box(0); -lean_inc(x_129); -x_149 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_31, x_132, x_1, x_148, x_36, x_129, x_147); -if (lean_obj_tag(x_149) == 0) -{ -lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; -x_150 = lean_ctor_get(x_149, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 1); -lean_inc(x_151); -lean_dec(x_149); -x_152 = lean_st_ref_get(x_129, x_151); -lean_dec(x_129); -x_153 = !lean_is_exclusive(x_152); -if (x_153 == 0) +x_146 = lean_box(0); +lean_inc(x_127); +x_147 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_29, x_130, x_1, x_146, x_34, x_127, x_145); +if (lean_obj_tag(x_147) == 0) { -lean_object* x_154; lean_object* x_155; -x_154 = lean_ctor_get(x_152, 0); -x_155 = lean_ctor_get(x_152, 1); -lean_ctor_set(x_152, 1, x_154); -lean_ctor_set(x_152, 0, x_150); -x_45 = x_152; -x_46 = x_155; -goto block_127; +lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; +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 = lean_st_ref_get(x_127, x_149); +lean_dec(x_127); +x_151 = !lean_is_exclusive(x_150); +if (x_151 == 0) +{ +lean_object* x_152; lean_object* x_153; +x_152 = lean_ctor_get(x_150, 0); +x_153 = lean_ctor_get(x_150, 1); +lean_ctor_set(x_150, 1, x_152); +lean_ctor_set(x_150, 0, x_148); +x_43 = x_150; +x_44 = x_153; +goto block_125; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_152, 0); -x_157 = lean_ctor_get(x_152, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_152); -x_158 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_158, 0, x_150); -lean_ctor_set(x_158, 1, x_156); -x_45 = x_158; -x_46 = x_157; -goto block_127; +lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_154 = lean_ctor_get(x_150, 0); +x_155 = lean_ctor_get(x_150, 1); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_150); +x_156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_156, 0, x_148); +lean_ctor_set(x_156, 1, x_154); +x_43 = x_156; +x_44 = x_155; +goto block_125; } } else { -uint8_t x_159; -lean_dec(x_129); -x_159 = !lean_is_exclusive(x_149); -if (x_159 == 0) +uint8_t x_157; +lean_dec(x_127); +x_157 = !lean_is_exclusive(x_147); +if (x_157 == 0) { -return x_149; +return x_147; } else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_160 = lean_ctor_get(x_149, 0); -x_161 = lean_ctor_get(x_149, 1); -lean_inc(x_161); -lean_inc(x_160); -lean_dec(x_149); -x_162 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_162, 0, x_160); -lean_ctor_set(x_162, 1, x_161); -return x_162; +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_147, 0); +x_159 = lean_ctor_get(x_147, 1); +lean_inc(x_159); +lean_inc(x_158); +lean_dec(x_147); +x_160 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_160, 0, x_158); +lean_ctor_set(x_160, 1, x_159); +return x_160; } } } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_163 = lean_ctor_get(x_140, 0); -x_164 = lean_ctor_get(x_140, 1); -x_165 = lean_ctor_get(x_140, 2); -x_166 = lean_ctor_get(x_140, 3); -x_167 = lean_ctor_get(x_140, 5); -x_168 = lean_ctor_get(x_140, 6); -x_169 = lean_ctor_get(x_140, 7); -lean_inc(x_169); -lean_inc(x_168); +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_161 = lean_ctor_get(x_138, 0); +x_162 = lean_ctor_get(x_138, 1); +x_163 = lean_ctor_get(x_138, 2); +x_164 = lean_ctor_get(x_138, 3); +x_165 = lean_ctor_get(x_138, 5); +x_166 = lean_ctor_get(x_138, 6); +x_167 = lean_ctor_get(x_138, 7); lean_inc(x_167); lean_inc(x_166); lean_inc(x_165); lean_inc(x_164); lean_inc(x_163); -lean_dec(x_140); -x_170 = l_Lean_Kernel_enableDiag(x_163, x_132); -x_171 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_171, 0, x_170); -lean_ctor_set(x_171, 1, x_164); -lean_ctor_set(x_171, 2, x_165); -lean_ctor_set(x_171, 3, x_166); -lean_ctor_set(x_171, 4, x_42); -lean_ctor_set(x_171, 5, x_167); -lean_ctor_set(x_171, 6, x_168); -lean_ctor_set(x_171, 7, x_169); -x_172 = lean_st_ref_set(x_129, x_171, x_141); -x_173 = lean_ctor_get(x_172, 1); -lean_inc(x_173); -lean_dec(x_172); -x_174 = lean_box(0); -lean_inc(x_129); -x_175 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_31, x_132, x_1, x_174, x_36, x_129, x_173); -if (lean_obj_tag(x_175) == 0) +lean_inc(x_162); +lean_inc(x_161); +lean_dec(x_138); +x_168 = l_Lean_Kernel_Environment_enableDiag(x_161, x_130); +x_169 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_162); +lean_ctor_set(x_169, 2, x_163); +lean_ctor_set(x_169, 3, x_164); +lean_ctor_set(x_169, 4, x_40); +lean_ctor_set(x_169, 5, x_165); +lean_ctor_set(x_169, 6, x_166); +lean_ctor_set(x_169, 7, x_167); +x_170 = lean_st_ref_set(x_127, x_169, x_139); +x_171 = lean_ctor_get(x_170, 1); +lean_inc(x_171); +lean_dec(x_170); +x_172 = lean_box(0); +lean_inc(x_127); +x_173 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_29, x_130, x_1, x_172, x_34, x_127, x_171); +if (lean_obj_tag(x_173) == 0) { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_176 = lean_ctor_get(x_175, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_175, 1); +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +lean_dec(x_173); +x_176 = lean_st_ref_get(x_127, x_175); +lean_dec(x_127); +x_177 = lean_ctor_get(x_176, 0); lean_inc(x_177); -lean_dec(x_175); -x_178 = lean_st_ref_get(x_129, x_177); -lean_dec(x_129); -x_179 = lean_ctor_get(x_178, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_178, 1); -lean_inc(x_180); -if (lean_is_exclusive(x_178)) { - lean_ctor_release(x_178, 0); - lean_ctor_release(x_178, 1); - x_181 = x_178; +x_178 = lean_ctor_get(x_176, 1); +lean_inc(x_178); +if (lean_is_exclusive(x_176)) { + lean_ctor_release(x_176, 0); + lean_ctor_release(x_176, 1); + x_179 = x_176; } else { - lean_dec_ref(x_178); - x_181 = lean_box(0); + lean_dec_ref(x_176); + x_179 = lean_box(0); } -if (lean_is_scalar(x_181)) { - x_182 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_179)) { + x_180 = lean_alloc_ctor(0, 2, 0); } else { - x_182 = x_181; + x_180 = x_179; } -lean_ctor_set(x_182, 0, x_176); -lean_ctor_set(x_182, 1, x_179); -x_45 = x_182; -x_46 = x_180; -goto block_127; +lean_ctor_set(x_180, 0, x_174); +lean_ctor_set(x_180, 1, x_177); +x_43 = x_180; +x_44 = x_178; +goto block_125; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -lean_dec(x_129); -x_183 = lean_ctor_get(x_175, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_175, 1); -lean_inc(x_184); -if (lean_is_exclusive(x_175)) { - lean_ctor_release(x_175, 0); - lean_ctor_release(x_175, 1); - x_185 = x_175; +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_127); +x_181 = lean_ctor_get(x_173, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_173, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_183 = x_173; } else { - lean_dec_ref(x_175); - x_185 = lean_box(0); + lean_dec_ref(x_173); + x_183 = lean_box(0); } -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_183)) { + x_184 = lean_alloc_ctor(1, 2, 0); } else { - x_186 = x_185; + x_184 = x_183; } -lean_ctor_set(x_186, 0, x_183); -lean_ctor_set(x_186, 1, x_184); -return x_186; +lean_ctor_set(x_184, 0, x_181); +lean_ctor_set(x_184, 1, x_182); +return x_184; } } } else { -lean_object* x_187; lean_object* x_188; -x_187 = lean_box(0); -lean_inc(x_129); -x_188 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_31, x_132, x_1, x_187, x_36, x_129, x_135); -if (lean_obj_tag(x_188) == 0) +lean_object* x_185; lean_object* x_186; +x_185 = lean_box(0); +lean_inc(x_127); +x_186 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_29, x_130, x_1, x_185, x_34, x_127, x_133); +if (lean_obj_tag(x_186) == 0) { -lean_object* x_189; lean_object* x_190; lean_object* x_191; uint8_t x_192; -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -x_190 = lean_ctor_get(x_188, 1); -lean_inc(x_190); -lean_dec(x_188); -x_191 = lean_st_ref_get(x_129, x_190); -lean_dec(x_129); -x_192 = !lean_is_exclusive(x_191); -if (x_192 == 0) +lean_object* x_187; lean_object* x_188; lean_object* x_189; uint8_t x_190; +x_187 = lean_ctor_get(x_186, 0); +lean_inc(x_187); +x_188 = lean_ctor_get(x_186, 1); +lean_inc(x_188); +lean_dec(x_186); +x_189 = lean_st_ref_get(x_127, x_188); +lean_dec(x_127); +x_190 = !lean_is_exclusive(x_189); +if (x_190 == 0) { -lean_object* x_193; lean_object* x_194; -x_193 = lean_ctor_get(x_191, 0); -x_194 = lean_ctor_get(x_191, 1); -lean_ctor_set(x_191, 1, x_193); -lean_ctor_set(x_191, 0, x_189); -x_45 = x_191; -x_46 = x_194; -goto block_127; +lean_object* x_191; lean_object* x_192; +x_191 = lean_ctor_get(x_189, 0); +x_192 = lean_ctor_get(x_189, 1); +lean_ctor_set(x_189, 1, x_191); +lean_ctor_set(x_189, 0, x_187); +x_43 = x_189; +x_44 = x_192; +goto block_125; } else { -lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_195 = lean_ctor_get(x_191, 0); -x_196 = lean_ctor_get(x_191, 1); -lean_inc(x_196); -lean_inc(x_195); -lean_dec(x_191); -x_197 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_197, 0, x_189); -lean_ctor_set(x_197, 1, x_195); -x_45 = x_197; -x_46 = x_196; -goto block_127; +lean_object* x_193; lean_object* x_194; lean_object* x_195; +x_193 = lean_ctor_get(x_189, 0); +x_194 = lean_ctor_get(x_189, 1); +lean_inc(x_194); +lean_inc(x_193); +lean_dec(x_189); +x_195 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_195, 0, x_187); +lean_ctor_set(x_195, 1, x_193); +x_43 = x_195; +x_44 = x_194; +goto block_125; } } else { -uint8_t x_198; -lean_dec(x_129); -x_198 = !lean_is_exclusive(x_188); -if (x_198 == 0) +uint8_t x_196; +lean_dec(x_127); +x_196 = !lean_is_exclusive(x_186); +if (x_196 == 0) { -return x_188; +return x_186; } else { -lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_199 = lean_ctor_get(x_188, 0); -x_200 = lean_ctor_get(x_188, 1); -lean_inc(x_200); -lean_inc(x_199); -lean_dec(x_188); -x_201 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_201, 0, x_199); -lean_ctor_set(x_201, 1, x_200); -return x_201; +lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_197 = lean_ctor_get(x_186, 0); +x_198 = lean_ctor_get(x_186, 1); +lean_inc(x_198); +lean_inc(x_197); +lean_dec(x_186); +x_199 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_199, 0, x_197); +lean_ctor_set(x_199, 1, x_198); +return x_199; } } } @@ -4548,391 +4526,391 @@ return x_201; } else { -uint8_t x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; uint8_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; uint8_t x_260; uint8_t x_261; -x_205 = lean_ctor_get_uint8(x_16, sizeof(void*)*2); +uint8_t x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; uint8_t x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; uint8_t x_258; uint8_t x_259; +x_203 = lean_ctor_get_uint8(x_16, sizeof(void*)*2); lean_dec(x_16); -x_206 = l_Lean_Elab_Command_mkState___closed__8; -x_207 = l_Lean_Elab_Command_mkState___closed__11; -x_208 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_208, 0, x_206); -lean_ctor_set(x_208, 1, x_207); -lean_ctor_set_uint8(x_208, sizeof(void*)*2, x_205); -x_209 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3; -x_210 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; -x_211 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_211, 0, x_21); -lean_ctor_set(x_211, 1, x_13); -lean_ctor_set(x_211, 2, x_15); -lean_ctor_set(x_211, 3, x_17); -lean_ctor_set(x_211, 4, x_209); -lean_ctor_set(x_211, 5, x_210); -lean_ctor_set(x_211, 6, x_208); -lean_ctor_set(x_211, 7, x_18); -x_251 = lean_st_mk_ref(x_211, x_10); -x_252 = lean_ctor_get(x_251, 0); -lean_inc(x_252); -x_253 = lean_ctor_get(x_251, 1); -lean_inc(x_253); -lean_dec(x_251); -x_254 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__5; -x_255 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_31, x_254); -x_256 = lean_st_ref_get(x_252, x_253); -x_257 = lean_ctor_get(x_256, 0); +x_204 = l_Lean_Elab_Command_mkState___closed__8; +x_205 = l_Lean_Elab_Command_mkState___closed__11; +x_206 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_206, 0, x_204); +lean_ctor_set(x_206, 1, x_205); +lean_ctor_set_uint8(x_206, sizeof(void*)*2, x_203); +x_207 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1; +x_208 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; +x_209 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_209, 0, x_19); +lean_ctor_set(x_209, 1, x_13); +lean_ctor_set(x_209, 2, x_15); +lean_ctor_set(x_209, 3, x_17); +lean_ctor_set(x_209, 4, x_207); +lean_ctor_set(x_209, 5, x_208); +lean_ctor_set(x_209, 6, x_206); +lean_ctor_set(x_209, 7, x_18); +x_249 = lean_st_mk_ref(x_209, x_10); +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_249, 1); +lean_inc(x_251); +lean_dec(x_249); +x_252 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3; +x_253 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_29, x_252); +x_254 = lean_st_ref_get(x_250, x_251); +x_255 = lean_ctor_get(x_254, 0); +lean_inc(x_255); +x_256 = lean_ctor_get(x_254, 1); +lean_inc(x_256); +lean_dec(x_254); +x_257 = lean_ctor_get(x_255, 0); lean_inc(x_257); -x_258 = lean_ctor_get(x_256, 1); -lean_inc(x_258); -lean_dec(x_256); -x_259 = lean_ctor_get(x_257, 0); -lean_inc(x_259); +lean_dec(x_255); +x_258 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_257); lean_dec(x_257); -x_260 = l_Lean_Kernel_isDiagnosticsEnabled(x_259); -lean_dec(x_259); -if (x_260 == 0) +if (x_258 == 0) { -if (x_255 == 0) +if (x_253 == 0) { -uint8_t x_304; -x_304 = 1; -x_261 = x_304; -goto block_303; +uint8_t x_302; +x_302 = 1; +x_259 = x_302; +goto block_301; } else { -x_261 = x_35; -goto block_303; +x_259 = x_33; +goto block_301; } } else { -if (x_255 == 0) +if (x_253 == 0) { -x_261 = x_35; -goto block_303; +x_259 = x_33; +goto block_301; } else { -uint8_t x_305; -x_305 = 1; -x_261 = x_305; -goto block_303; +uint8_t x_303; +x_303 = 1; +x_259 = x_303; +goto block_301; } } -block_250: +block_248: { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; uint64_t x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -x_214 = lean_ctor_get(x_212, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_212, 1); +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; uint8_t x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; uint64_t x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; +x_212 = lean_ctor_get(x_210, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_210, 1); +lean_inc(x_213); +lean_dec(x_210); +x_214 = lean_st_ref_take(x_3, x_211); +x_215 = lean_ctor_get(x_214, 0); lean_inc(x_215); -lean_dec(x_212); -x_216 = lean_st_ref_take(x_3, x_213); -x_217 = lean_ctor_get(x_216, 0); +x_216 = lean_ctor_get(x_214, 1); +lean_inc(x_216); +lean_dec(x_214); +x_217 = lean_ctor_get(x_213, 0); lean_inc(x_217); -x_218 = lean_ctor_get(x_216, 1); +x_218 = lean_ctor_get(x_213, 1); lean_inc(x_218); -lean_dec(x_216); -x_219 = lean_ctor_get(x_215, 0); +x_219 = lean_ctor_get(x_213, 2); lean_inc(x_219); -x_220 = lean_ctor_get(x_215, 1); +x_220 = lean_ctor_get(x_213, 3); lean_inc(x_220); -x_221 = lean_ctor_get(x_215, 2); +x_221 = lean_ctor_get(x_213, 5); lean_inc(x_221); -x_222 = lean_ctor_get(x_215, 3); +x_222 = lean_ctor_get(x_213, 6); lean_inc(x_222); -x_223 = lean_ctor_get(x_215, 5); +x_223 = lean_ctor_get(x_213, 7); lean_inc(x_223); -x_224 = lean_ctor_get(x_215, 6); +lean_dec(x_213); +x_224 = lean_ctor_get(x_215, 1); lean_inc(x_224); -x_225 = lean_ctor_get(x_215, 7); +x_225 = lean_ctor_get(x_215, 2); lean_inc(x_225); -lean_dec(x_215); -x_226 = lean_ctor_get(x_217, 1); +x_226 = lean_ctor_get(x_215, 4); lean_inc(x_226); -x_227 = lean_ctor_get(x_217, 2); +x_227 = lean_ctor_get(x_215, 6); lean_inc(x_227); -x_228 = lean_ctor_get(x_217, 4); +x_228 = lean_ctor_get(x_215, 7); lean_inc(x_228); -x_229 = lean_ctor_get(x_217, 6); -lean_inc(x_229); -x_230 = lean_ctor_get(x_217, 7); -lean_inc(x_230); -if (lean_is_exclusive(x_217)) { - lean_ctor_release(x_217, 0); - lean_ctor_release(x_217, 1); - lean_ctor_release(x_217, 2); - lean_ctor_release(x_217, 3); - lean_ctor_release(x_217, 4); - lean_ctor_release(x_217, 5); - lean_ctor_release(x_217, 6); - lean_ctor_release(x_217, 7); - lean_ctor_release(x_217, 8); - x_231 = x_217; -} else { - lean_dec_ref(x_217); - x_231 = lean_box(0); -} -x_232 = l_Lean_MessageLog_append(x_226, x_223); -x_233 = lean_ctor_get_uint8(x_229, sizeof(void*)*2); -x_234 = lean_ctor_get(x_229, 0); -lean_inc(x_234); -x_235 = lean_ctor_get(x_229, 1); +if (lean_is_exclusive(x_215)) { + lean_ctor_release(x_215, 0); + lean_ctor_release(x_215, 1); + lean_ctor_release(x_215, 2); + lean_ctor_release(x_215, 3); + lean_ctor_release(x_215, 4); + lean_ctor_release(x_215, 5); + lean_ctor_release(x_215, 6); + lean_ctor_release(x_215, 7); + lean_ctor_release(x_215, 8); + x_229 = x_215; +} else { + lean_dec_ref(x_215); + x_229 = lean_box(0); +} +x_230 = l_Lean_MessageLog_append(x_224, x_221); +x_231 = lean_ctor_get_uint8(x_227, sizeof(void*)*2); +x_232 = lean_ctor_get(x_227, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_227, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_227)) { + lean_ctor_release(x_227, 0); + lean_ctor_release(x_227, 1); + x_234 = x_227; +} else { + lean_dec_ref(x_227); + x_234 = lean_box(0); +} +x_235 = lean_ctor_get(x_222, 1); lean_inc(x_235); -if (lean_is_exclusive(x_229)) { - lean_ctor_release(x_229, 0); - lean_ctor_release(x_229, 1); - x_236 = x_229; +lean_dec(x_222); +x_236 = l_Lean_PersistentArray_append___rarg(x_233, x_235); +lean_dec(x_235); +if (lean_is_scalar(x_234)) { + x_237 = lean_alloc_ctor(0, 2, 1); } else { - lean_dec_ref(x_229); - x_236 = lean_box(0); + x_237 = x_234; } -x_237 = lean_ctor_get(x_224, 1); -lean_inc(x_237); -lean_dec(x_224); -x_238 = l_Lean_PersistentArray_append___rarg(x_235, x_237); -lean_dec(x_237); -if (lean_is_scalar(x_236)) { - x_239 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_237, 0, x_232); +lean_ctor_set(x_237, 1, x_236); +lean_ctor_set_uint8(x_237, sizeof(void*)*2, x_231); +x_238 = lean_ctor_get_uint64(x_228, sizeof(void*)*1); +if (lean_is_exclusive(x_228)) { + lean_ctor_release(x_228, 0); + x_239 = x_228; } else { - x_239 = x_236; + lean_dec_ref(x_228); + x_239 = lean_box(0); } -lean_ctor_set(x_239, 0, x_234); -lean_ctor_set(x_239, 1, x_238); -lean_ctor_set_uint8(x_239, sizeof(void*)*2, x_233); -x_240 = lean_ctor_get_uint64(x_230, sizeof(void*)*1); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - x_241 = x_230; +x_240 = lean_ctor_get(x_220, 0); +lean_inc(x_240); +lean_dec(x_220); +x_241 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_26, x_240); +if (lean_is_scalar(x_239)) { + x_242 = lean_alloc_ctor(0, 1, 8); } else { - lean_dec_ref(x_230); - x_241 = lean_box(0); + x_242 = x_239; } -x_242 = lean_ctor_get(x_222, 0); -lean_inc(x_242); -lean_dec(x_222); -x_243 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_28, x_242); -if (lean_is_scalar(x_241)) { - x_244 = lean_alloc_ctor(0, 1, 8); -} else { - x_244 = x_241; -} -lean_ctor_set(x_244, 0, x_243); -lean_ctor_set_uint64(x_244, sizeof(void*)*1, x_240); -if (lean_is_scalar(x_231)) { - x_245 = lean_alloc_ctor(0, 9, 0); -} else { - x_245 = x_231; -} -lean_ctor_set(x_245, 0, x_219); -lean_ctor_set(x_245, 1, x_232); -lean_ctor_set(x_245, 2, x_227); -lean_ctor_set(x_245, 3, x_220); -lean_ctor_set(x_245, 4, x_228); -lean_ctor_set(x_245, 5, x_221); -lean_ctor_set(x_245, 6, x_239); -lean_ctor_set(x_245, 7, x_244); -lean_ctor_set(x_245, 8, x_225); -x_246 = lean_st_ref_set(x_3, x_245, x_218); -x_247 = lean_ctor_get(x_246, 1); -lean_inc(x_247); -if (lean_is_exclusive(x_246)) { - lean_ctor_release(x_246, 0); - lean_ctor_release(x_246, 1); - x_248 = x_246; +lean_ctor_set(x_242, 0, x_241); +lean_ctor_set_uint64(x_242, sizeof(void*)*1, x_238); +if (lean_is_scalar(x_229)) { + x_243 = lean_alloc_ctor(0, 9, 0); +} else { + x_243 = x_229; +} +lean_ctor_set(x_243, 0, x_217); +lean_ctor_set(x_243, 1, x_230); +lean_ctor_set(x_243, 2, x_225); +lean_ctor_set(x_243, 3, x_218); +lean_ctor_set(x_243, 4, x_226); +lean_ctor_set(x_243, 5, x_219); +lean_ctor_set(x_243, 6, x_237); +lean_ctor_set(x_243, 7, x_242); +lean_ctor_set(x_243, 8, x_223); +x_244 = lean_st_ref_set(x_3, x_243, x_216); +x_245 = lean_ctor_get(x_244, 1); +lean_inc(x_245); +if (lean_is_exclusive(x_244)) { + lean_ctor_release(x_244, 0); + lean_ctor_release(x_244, 1); + x_246 = x_244; } else { - lean_dec_ref(x_246); - x_248 = lean_box(0); + lean_dec_ref(x_244); + x_246 = lean_box(0); } -if (lean_is_scalar(x_248)) { - x_249 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_246)) { + x_247 = lean_alloc_ctor(0, 2, 0); } else { - x_249 = x_248; + x_247 = x_246; } -lean_ctor_set(x_249, 0, x_214); -lean_ctor_set(x_249, 1, x_247); -return x_249; +lean_ctor_set(x_247, 0, x_212); +lean_ctor_set(x_247, 1, x_245); +return x_247; } -block_303: +block_301: { -if (x_261 == 0) +if (x_259 == 0) { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; -x_262 = lean_st_ref_take(x_252, x_258); -x_263 = lean_ctor_get(x_262, 0); +lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +x_260 = lean_st_ref_take(x_250, x_256); +x_261 = lean_ctor_get(x_260, 0); +lean_inc(x_261); +x_262 = lean_ctor_get(x_260, 1); +lean_inc(x_262); +lean_dec(x_260); +x_263 = lean_ctor_get(x_261, 0); lean_inc(x_263); -x_264 = lean_ctor_get(x_262, 1); +x_264 = lean_ctor_get(x_261, 1); lean_inc(x_264); -lean_dec(x_262); -x_265 = lean_ctor_get(x_263, 0); +x_265 = lean_ctor_get(x_261, 2); lean_inc(x_265); -x_266 = lean_ctor_get(x_263, 1); +x_266 = lean_ctor_get(x_261, 3); lean_inc(x_266); -x_267 = lean_ctor_get(x_263, 2); +x_267 = lean_ctor_get(x_261, 5); lean_inc(x_267); -x_268 = lean_ctor_get(x_263, 3); +x_268 = lean_ctor_get(x_261, 6); lean_inc(x_268); -x_269 = lean_ctor_get(x_263, 5); +x_269 = lean_ctor_get(x_261, 7); lean_inc(x_269); -x_270 = lean_ctor_get(x_263, 6); -lean_inc(x_270); -x_271 = lean_ctor_get(x_263, 7); -lean_inc(x_271); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - lean_ctor_release(x_263, 2); - lean_ctor_release(x_263, 3); - lean_ctor_release(x_263, 4); - lean_ctor_release(x_263, 5); - lean_ctor_release(x_263, 6); - lean_ctor_release(x_263, 7); - x_272 = x_263; -} else { - lean_dec_ref(x_263); - x_272 = lean_box(0); +if (lean_is_exclusive(x_261)) { + lean_ctor_release(x_261, 0); + lean_ctor_release(x_261, 1); + lean_ctor_release(x_261, 2); + lean_ctor_release(x_261, 3); + lean_ctor_release(x_261, 4); + lean_ctor_release(x_261, 5); + lean_ctor_release(x_261, 6); + lean_ctor_release(x_261, 7); + x_270 = x_261; +} else { + lean_dec_ref(x_261); + x_270 = lean_box(0); } -x_273 = l_Lean_Kernel_enableDiag(x_265, x_255); -if (lean_is_scalar(x_272)) { - x_274 = lean_alloc_ctor(0, 8, 0); -} else { - x_274 = x_272; -} -lean_ctor_set(x_274, 0, x_273); -lean_ctor_set(x_274, 1, x_266); -lean_ctor_set(x_274, 2, x_267); -lean_ctor_set(x_274, 3, x_268); -lean_ctor_set(x_274, 4, x_209); -lean_ctor_set(x_274, 5, x_269); -lean_ctor_set(x_274, 6, x_270); -lean_ctor_set(x_274, 7, x_271); -x_275 = lean_st_ref_set(x_252, x_274, x_264); -x_276 = lean_ctor_get(x_275, 1); -lean_inc(x_276); -lean_dec(x_275); -x_277 = lean_box(0); -lean_inc(x_252); -x_278 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_31, x_255, x_1, x_277, x_36, x_252, x_276); -if (lean_obj_tag(x_278) == 0) -{ -lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; -x_279 = lean_ctor_get(x_278, 0); -lean_inc(x_279); -x_280 = lean_ctor_get(x_278, 1); +x_271 = l_Lean_Kernel_Environment_enableDiag(x_263, x_253); +if (lean_is_scalar(x_270)) { + x_272 = lean_alloc_ctor(0, 8, 0); +} else { + x_272 = x_270; +} +lean_ctor_set(x_272, 0, x_271); +lean_ctor_set(x_272, 1, x_264); +lean_ctor_set(x_272, 2, x_265); +lean_ctor_set(x_272, 3, x_266); +lean_ctor_set(x_272, 4, x_207); +lean_ctor_set(x_272, 5, x_267); +lean_ctor_set(x_272, 6, x_268); +lean_ctor_set(x_272, 7, x_269); +x_273 = lean_st_ref_set(x_250, x_272, x_262); +x_274 = lean_ctor_get(x_273, 1); +lean_inc(x_274); +lean_dec(x_273); +x_275 = lean_box(0); +lean_inc(x_250); +x_276 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_29, x_253, x_1, x_275, x_34, x_250, x_274); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_277 = lean_ctor_get(x_276, 0); +lean_inc(x_277); +x_278 = lean_ctor_get(x_276, 1); +lean_inc(x_278); +lean_dec(x_276); +x_279 = lean_st_ref_get(x_250, x_278); +lean_dec(x_250); +x_280 = lean_ctor_get(x_279, 0); lean_inc(x_280); -lean_dec(x_278); -x_281 = lean_st_ref_get(x_252, x_280); -lean_dec(x_252); -x_282 = lean_ctor_get(x_281, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_281, 1); -lean_inc(x_283); -if (lean_is_exclusive(x_281)) { - lean_ctor_release(x_281, 0); - lean_ctor_release(x_281, 1); - x_284 = x_281; -} else { - lean_dec_ref(x_281); - x_284 = lean_box(0); +x_281 = lean_ctor_get(x_279, 1); +lean_inc(x_281); +if (lean_is_exclusive(x_279)) { + lean_ctor_release(x_279, 0); + lean_ctor_release(x_279, 1); + x_282 = x_279; +} else { + lean_dec_ref(x_279); + x_282 = lean_box(0); } -if (lean_is_scalar(x_284)) { - x_285 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_282)) { + x_283 = lean_alloc_ctor(0, 2, 0); } else { - x_285 = x_284; + x_283 = x_282; } -lean_ctor_set(x_285, 0, x_279); -lean_ctor_set(x_285, 1, x_282); -x_212 = x_285; -x_213 = x_283; -goto block_250; +lean_ctor_set(x_283, 0, x_277); +lean_ctor_set(x_283, 1, x_280); +x_210 = x_283; +x_211 = x_281; +goto block_248; } else { -lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; -lean_dec(x_252); -x_286 = lean_ctor_get(x_278, 0); -lean_inc(x_286); -x_287 = lean_ctor_get(x_278, 1); -lean_inc(x_287); -if (lean_is_exclusive(x_278)) { - lean_ctor_release(x_278, 0); - lean_ctor_release(x_278, 1); - x_288 = x_278; +lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; +lean_dec(x_250); +x_284 = lean_ctor_get(x_276, 0); +lean_inc(x_284); +x_285 = lean_ctor_get(x_276, 1); +lean_inc(x_285); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + lean_ctor_release(x_276, 1); + x_286 = x_276; } else { - lean_dec_ref(x_278); - x_288 = lean_box(0); + lean_dec_ref(x_276); + x_286 = lean_box(0); } -if (lean_is_scalar(x_288)) { - x_289 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_286)) { + x_287 = lean_alloc_ctor(1, 2, 0); } else { - x_289 = x_288; + x_287 = x_286; } -lean_ctor_set(x_289, 0, x_286); -lean_ctor_set(x_289, 1, x_287); -return x_289; +lean_ctor_set(x_287, 0, x_284); +lean_ctor_set(x_287, 1, x_285); +return x_287; } } else { -lean_object* x_290; lean_object* x_291; -x_290 = lean_box(0); -lean_inc(x_252); -x_291 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_31, x_255, x_1, x_290, x_36, x_252, x_258); -if (lean_obj_tag(x_291) == 0) +lean_object* x_288; lean_object* x_289; +x_288 = lean_box(0); +lean_inc(x_250); +x_289 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_29, x_253, x_1, x_288, x_34, x_250, x_256); +if (lean_obj_tag(x_289) == 0) { -lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_292 = lean_ctor_get(x_291, 0); -lean_inc(x_292); -x_293 = lean_ctor_get(x_291, 1); +lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +x_290 = lean_ctor_get(x_289, 0); +lean_inc(x_290); +x_291 = lean_ctor_get(x_289, 1); +lean_inc(x_291); +lean_dec(x_289); +x_292 = lean_st_ref_get(x_250, x_291); +lean_dec(x_250); +x_293 = lean_ctor_get(x_292, 0); lean_inc(x_293); -lean_dec(x_291); -x_294 = lean_st_ref_get(x_252, x_293); -lean_dec(x_252); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -x_296 = lean_ctor_get(x_294, 1); -lean_inc(x_296); -if (lean_is_exclusive(x_294)) { - lean_ctor_release(x_294, 0); - lean_ctor_release(x_294, 1); - x_297 = x_294; +x_294 = lean_ctor_get(x_292, 1); +lean_inc(x_294); +if (lean_is_exclusive(x_292)) { + lean_ctor_release(x_292, 0); + lean_ctor_release(x_292, 1); + x_295 = x_292; } else { - lean_dec_ref(x_294); - x_297 = lean_box(0); + lean_dec_ref(x_292); + x_295 = lean_box(0); } -if (lean_is_scalar(x_297)) { - x_298 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_295)) { + x_296 = lean_alloc_ctor(0, 2, 0); } else { - x_298 = x_297; + x_296 = x_295; } -lean_ctor_set(x_298, 0, x_292); -lean_ctor_set(x_298, 1, x_295); -x_212 = x_298; -x_213 = x_296; -goto block_250; +lean_ctor_set(x_296, 0, x_290); +lean_ctor_set(x_296, 1, x_293); +x_210 = x_296; +x_211 = x_294; +goto block_248; } else { -lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; -lean_dec(x_252); -x_299 = lean_ctor_get(x_291, 0); -lean_inc(x_299); -x_300 = lean_ctor_get(x_291, 1); -lean_inc(x_300); -if (lean_is_exclusive(x_291)) { - lean_ctor_release(x_291, 0); - lean_ctor_release(x_291, 1); - x_301 = x_291; +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +lean_dec(x_250); +x_297 = lean_ctor_get(x_289, 0); +lean_inc(x_297); +x_298 = lean_ctor_get(x_289, 1); +lean_inc(x_298); +if (lean_is_exclusive(x_289)) { + lean_ctor_release(x_289, 0); + lean_ctor_release(x_289, 1); + x_299 = x_289; } else { - lean_dec_ref(x_291); - x_301 = lean_box(0); + lean_dec_ref(x_289); + x_299 = lean_box(0); } -if (lean_is_scalar(x_301)) { - x_302 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_299)) { + x_300 = lean_alloc_ctor(1, 2, 0); } else { - x_302 = x_301; + x_300 = x_299; } -lean_ctor_set(x_302, 0, x_299); -lean_ctor_set(x_302, 1, x_300); -return x_302; +lean_ctor_set(x_300, 0, x_297); +lean_ctor_set(x_300, 1, x_298); +return x_300; } } } @@ -21395,7 +21373,8 @@ x_5 = lean_ctor_get(x_3, 0); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); lean_dec(x_5); -x_7 = lean_environment_main_module(x_6); +x_7 = l_Lean_Environment_mainModule(x_6); +lean_dec(x_6); lean_ctor_set(x_3, 0, x_7); return x_3; } @@ -21410,7 +21389,8 @@ lean_dec(x_3); x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); -x_11 = lean_environment_main_module(x_10); +x_11 = l_Lean_Environment_mainModule(x_10); +lean_dec(x_10); x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_9); @@ -26340,7 +26320,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 3); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_8); +x_39 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_40 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_40, 0, x_22); lean_ctor_set(x_40, 1, x_39); @@ -26521,7 +26502,8 @@ lean_dec(x_34); x_89 = lean_ctor_get(x_87, 3); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_8); +x_90 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_91 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_91, 0, x_22); lean_ctor_set(x_91, 1, x_90); @@ -26944,7 +26926,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 3); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_8); +x_39 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_40 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_40, 0, x_22); lean_ctor_set(x_40, 1, x_39); @@ -27125,7 +27108,8 @@ lean_dec(x_34); x_89 = lean_ctor_get(x_87, 3); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_8); +x_90 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_91 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_91, 0, x_22); lean_ctor_set(x_91, 1, x_90); @@ -33720,7 +33704,7 @@ 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; lean_object* x_37; lean_object* x_38; lean_object* x_65; lean_object* x_66; lean_object* x_93; x_30 = lean_ctor_get(x_26, 1); -x_31 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; +x_31 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; lean_ctor_set(x_26, 1, x_31); x_32 = lean_st_ref_set(x_3, x_26, x_27); x_33 = lean_ctor_get(x_32, 1); @@ -34073,7 +34057,7 @@ lean_inc(x_108); lean_inc(x_107); lean_inc(x_106); lean_dec(x_26); -x_115 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; +x_115 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; x_116 = lean_alloc_ctor(0, 9, 0); lean_ctor_set(x_116, 0, x_106); lean_ctor_set(x_116, 1, x_115); @@ -40778,7 +40762,7 @@ x_35 = l_Lean_Elab_Command_mkState___closed__8; x_36 = l_Lean_Elab_Command_mkState___closed__11; lean_ctor_set(x_26, 1, x_36); lean_ctor_set(x_26, 0, x_35); -x_37 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; +x_37 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; x_38 = l_Lean_Elab_Command_mkState___closed__13; lean_inc(x_14); x_39 = lean_alloc_ctor(0, 9, 0); @@ -41331,7 +41315,7 @@ x_187 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_187, 0, x_185); lean_ctor_set(x_187, 1, x_186); lean_ctor_set_uint8(x_187, sizeof(void*)*2, x_184); -x_188 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; +x_188 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; x_189 = l_Lean_Elab_Command_mkState___closed__13; lean_inc(x_14); x_190 = lean_alloc_ctor(0, 9, 0); @@ -41682,7 +41666,7 @@ if (lean_is_scalar(x_284)) { lean_ctor_set(x_287, 0, x_285); lean_ctor_set(x_287, 1, x_286); lean_ctor_set_uint8(x_287, sizeof(void*)*2, x_283); -x_288 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; +x_288 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; x_289 = l_Lean_Elab_Command_mkState___closed__13; lean_inc(x_264); x_290 = lean_alloc_ctor(0, 9, 0); @@ -43953,10 +43937,6 @@ l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2 = lean_mark_persistent(l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2); l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3 = _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3(); lean_mark_persistent(l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3); -l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4 = _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4(); -lean_mark_persistent(l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4); -l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__5 = _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__5(); -lean_mark_persistent(l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__5); l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__1 = _init_l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__1); l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__2 = _init_l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/ComputedFields.c b/stage0/stdlib/Lean/Elab/ComputedFields.c index a1eacea986be..bc9c2c329713 100644 --- a/stage0/stdlib/Lean/Elab/ComputedFields.c +++ b/stage0/stdlib/Lean/Elab/ComputedFields.c @@ -68,7 +68,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_ComputedFields_overrideCasesOn___lambda__1( static lean_object* l_Lean_Elab_ComputedFields_getComputedFieldValue___lambda__1___closed__4; static lean_object* l_Lean_Elab_ComputedFields_initFn____x40_Lean_Elab_ComputedFields___hyg_5____closed__1; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_ComputedFields_mkImplType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_ComputedFields_overrideCasesOn___spec__6___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_ComputedFields_overrideConstructors___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4220,8 +4220,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -4261,8 +4260,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; diff --git a/stage0/stdlib/Lean/Elab/DeclModifiers.c b/stage0/stdlib/Lean/Elab/DeclModifiers.c index 1f95560e2047..78a45c2ba55c 100644 --- a/stage0/stdlib/Lean/Elab/DeclModifiers.c +++ b/stage0/stdlib/Lean/Elab/DeclModifiers.c @@ -271,6 +271,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___rarg___lambda__7( static lean_object* l_Lean_Elab_instToFormatModifiers___closed__13; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___rarg___lambda__5___closed__2; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_applyVisibility___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Modifiers_filterAttrs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_instToFormatModifiers___closed__29; @@ -585,7 +586,6 @@ lean_closure_set(x_9, 4, x_5); lean_closure_set(x_9, 5, x_6); lean_closure_set(x_9, 6, x_7); lean_inc(x_1); -lean_inc(x_3); x_10 = l_Lean_mkPrivateName(x_3, x_1); x_11 = l_Lean_Environment_contains(x_3, x_10); if (x_11 == 0) @@ -3325,7 +3325,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_2, 0); lean_inc(x_16); lean_inc(x_15); -x_17 = lean_alloc_closure((void*)(l_Lean_Elab_applyVisibility___rarg___lambda__4), 7, 6); +x_17 = lean_alloc_closure((void*)(l_Lean_Elab_applyVisibility___rarg___lambda__4___boxed), 7, 6); lean_closure_set(x_17, 0, x_6); lean_closure_set(x_17, 1, x_1); lean_closure_set(x_17, 2, x_2); @@ -3364,6 +3364,15 @@ lean_dec(x_5); return x_6; } } +LEAN_EXPORT lean_object* l_Lean_Elab_applyVisibility___rarg___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Elab_applyVisibility___rarg___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_7); +return x_8; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_applyVisibility___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { diff --git a/stage0/stdlib/Lean/Elab/DeclNameGen.c b/stage0/stdlib/Lean/Elab/DeclNameGen.c index 747db535e6fc..426cbc5ba8a4 100644 --- a/stage0/stdlib/Lean/Elab/DeclNameGen.c +++ b/stage0/stdlib/Lean/Elab/DeclNameGen.c @@ -62,7 +62,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_ LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_NameGen_winnowExpr_visit___spec__9___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 uint8_t l_Lean_Elab_Command_NameGen_mkBaseNameWithSuffix___lambda__1(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_NameGen_mkBaseNameWithSuffix_x27___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_NameGen_mkBaseNameAux_visit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -166,7 +166,7 @@ lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_NameGen_winnowExpr_visit___spec__9___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_NameGen_winnowExpr_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_NameGen_visitNamespace___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -274,10 +274,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_ { lean_object* x_11; lean_object* x_12; x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -lean_dec(x_1); lean_inc(x_2); -x_12 = lean_environment_find(x_2, x_11); +x_12 = l_Lean_Environment_find_x3f(x_2, x_11); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; @@ -419,6 +417,7 @@ lean_object* x_27; lean_object* x_28; lean_free_object(x_10); x_27 = lean_box(0); x_28 = l___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_NameGen_getParentProjArg___lambda__2(x_19, x_14, x_1, x_9, x_27, x_2, x_3, x_4, x_5, x_13); +lean_dec(x_19); return x_28; } } @@ -481,6 +480,7 @@ else lean_object* x_46; lean_object* x_47; x_46 = lean_box(0); x_47 = l___private_Lean_Elab_DeclNameGen_0__Lean_Elab_Command_NameGen_getParentProjArg___lambda__2(x_37, x_31, x_1, x_9, x_46, x_2, x_3, x_4, x_5, x_30); +lean_dec(x_37); return x_47; } } @@ -534,6 +534,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); +lean_dec(x_1); return x_11; } } @@ -5996,7 +5997,8 @@ if (lean_is_exclusive(x_20)) { x_24 = lean_ctor_get(x_21, 0); lean_inc(x_24); lean_dec(x_21); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Lean_Name_getRoot(x_25); lean_dec(x_25); x_27 = lean_ctor_get(x_17, 1); @@ -6241,7 +6243,8 @@ if (lean_is_exclusive(x_77)) { x_81 = lean_ctor_get(x_78, 0); lean_inc(x_81); lean_dec(x_78); -x_82 = lean_environment_main_module(x_81); +x_82 = l_Lean_Environment_mainModule(x_81); +lean_dec(x_81); x_83 = l_Lean_Name_getRoot(x_82); lean_dec(x_82); x_84 = lean_ctor_get(x_74, 1); @@ -6998,7 +7001,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -7180,7 +7184,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); diff --git a/stage0/stdlib/Lean/Elab/Declaration.c b/stage0/stdlib/Lean/Elab/Declaration.c index c7e383470601..701f004918e9 100644 --- a/stage0/stdlib/Lean/Elab/Declaration.c +++ b/stage0/stdlib/Lean/Elab/Declaration.c @@ -12367,6 +12367,7 @@ x_34 = lean_ctor_get(x_32, 0); lean_inc(x_34); lean_dec(x_32); x_35 = l_Lean_mkPrivateName(x_34, x_23); +lean_dec(x_34); x_36 = lean_box(0); x_37 = l_Lean_Elab_Command_elabInitialize___lambda__2(x_16, x_2, x_3, x_11, x_4, x_35, x_36, x_12, x_13, x_33); return x_37; diff --git a/stage0/stdlib/Lean/Elab/DefView.c b/stage0/stdlib/Lean/Elab/DefView.c index 4b630fbf1c2d..a2190cac2d28 100644 --- a/stage0/stdlib/Lean/Elab/DefView.c +++ b/stage0/stdlib/Lean/Elab/DefView.c @@ -184,7 +184,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Comma static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__6; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_2783____closed__8; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkDefViewOfInstance___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_mkDefViewOfInstance___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1955,7 +1955,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 3); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_8); +x_39 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_40 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_40, 0, x_22); lean_ctor_set(x_40, 1, x_39); @@ -2136,7 +2137,8 @@ lean_dec(x_34); x_89 = lean_ctor_get(x_87, 3); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_8); +x_90 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_91 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_91, 0, x_22); lean_ctor_set(x_91, 1, x_90); @@ -2559,7 +2561,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 3); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_8); +x_39 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_40 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_40, 0, x_22); lean_ctor_set(x_40, 1, x_39); @@ -2740,7 +2743,8 @@ lean_dec(x_34); x_89 = lean_ctor_get(x_87, 3); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_8); +x_90 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_91 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_91, 0, x_22); lean_ctor_set(x_91, 1, x_90); diff --git a/stage0/stdlib/Lean/Elab/Deriving/BEq.c b/stage0/stdlib/Lean/Elab/Deriving/BEq.c index 6fdf5a7d1827..d23a6ffc9208 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/BEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/BEq.c @@ -73,7 +73,7 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__7; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___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*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3927____closed__14; static lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3927____closed__16; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; @@ -240,7 +240,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__18; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__6___lambda__1___closed__21; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_BEq_mkBEqInstanceHandler___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_allM___at_Lean_Elab_Deriving_BEq_mkBEqInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -719,7 +719,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 0); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_38); +x_39 = l_Lean_Environment_mainModule(x_38); +lean_dec(x_38); x_40 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__4; x_41 = l_Lean_addMacroScope(x_39, x_40, x_33); x_42 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__3; @@ -821,7 +822,8 @@ lean_dec(x_34); x_80 = lean_ctor_get(x_78, 0); lean_inc(x_80); lean_dec(x_78); -x_81 = lean_environment_main_module(x_80); +x_81 = l_Lean_Environment_mainModule(x_80); +lean_dec(x_80); x_82 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__4; x_83 = l_Lean_addMacroScope(x_81, x_82, x_33); x_84 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__3; @@ -908,7 +910,8 @@ if (lean_is_exclusive(x_109)) { x_113 = lean_ctor_get(x_110, 0); lean_inc(x_113); lean_dec(x_110); -x_114 = lean_environment_main_module(x_113); +x_114 = l_Lean_Environment_mainModule(x_113); +lean_dec(x_113); x_115 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__4; x_116 = l_Lean_addMacroScope(x_114, x_115, x_108); x_117 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__3; @@ -1021,7 +1024,8 @@ if (lean_is_exclusive(x_151)) { x_155 = lean_ctor_get(x_152, 0); lean_inc(x_155); lean_dec(x_152); -x_156 = lean_environment_main_module(x_155); +x_156 = l_Lean_Environment_mainModule(x_155); +lean_dec(x_155); x_157 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__4; x_158 = l_Lean_addMacroScope(x_156, x_157, x_150); x_159 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__3; @@ -2401,7 +2405,8 @@ x_102 = lean_ctor_get(x_100, 0); x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); lean_dec(x_102); -x_104 = lean_environment_main_module(x_103); +x_104 = l_Lean_Environment_mainModule(x_103); +lean_dec(x_103); x_105 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__6___lambda__1___closed__9; lean_inc(x_98); x_106 = lean_alloc_ctor(2, 2, 0); @@ -2574,7 +2579,8 @@ lean_dec(x_100); x_178 = lean_ctor_get(x_176, 0); lean_inc(x_178); lean_dec(x_176); -x_179 = lean_environment_main_module(x_178); +x_179 = l_Lean_Environment_mainModule(x_178); +lean_dec(x_178); x_180 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__6___lambda__1___closed__9; lean_inc(x_98); x_181 = lean_alloc_ctor(2, 2, 0); @@ -4225,7 +4231,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 0); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_38); +x_39 = l_Lean_Environment_mainModule(x_38); +lean_dec(x_38); x_40 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__8___lambda__1___closed__8; x_41 = l_Lean_addMacroScope(x_39, x_40, x_33); x_42 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__8___lambda__1___closed__7; @@ -4929,7 +4936,8 @@ lean_dec(x_34); x_285 = lean_ctor_get(x_283, 0); lean_inc(x_285); lean_dec(x_283); -x_286 = lean_environment_main_module(x_285); +x_286 = l_Lean_Environment_mainModule(x_285); +lean_dec(x_285); x_287 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__8___lambda__1___closed__8; x_288 = l_Lean_addMacroScope(x_286, x_287, x_33); x_289 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__8___lambda__1___closed__7; @@ -5210,7 +5218,8 @@ if (lean_is_exclusive(x_368)) { x_372 = lean_ctor_get(x_369, 0); lean_inc(x_372); lean_dec(x_369); -x_373 = lean_environment_main_module(x_372); +x_373 = l_Lean_Environment_mainModule(x_372); +lean_dec(x_372); x_374 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__8___lambda__1___closed__8; x_375 = l_Lean_addMacroScope(x_373, x_374, x_367); x_376 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__8___lambda__1___closed__7; @@ -6613,7 +6622,8 @@ x_21 = lean_ctor_get(x_19, 0); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec(x_21); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; x_25 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__14; lean_inc(x_17); @@ -6706,7 +6716,8 @@ lean_dec(x_19); x_65 = lean_ctor_get(x_63, 0); lean_inc(x_65); lean_dec(x_63); -x_66 = lean_environment_main_module(x_65); +x_66 = l_Lean_Environment_mainModule(x_65); +lean_dec(x_65); x_67 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; x_68 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__14; lean_inc(x_17); @@ -6812,7 +6823,8 @@ x_114 = lean_ctor_get(x_112, 0); x_115 = lean_ctor_get(x_114, 0); lean_inc(x_115); lean_dec(x_114); -x_116 = lean_environment_main_module(x_115); +x_116 = l_Lean_Environment_mainModule(x_115); +lean_dec(x_115); x_117 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; x_118 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__14; lean_inc(x_110); @@ -6915,7 +6927,8 @@ lean_dec(x_112); x_163 = lean_ctor_get(x_161, 0); lean_inc(x_163); lean_dec(x_161); -x_164 = lean_environment_main_module(x_163); +x_164 = l_Lean_Environment_mainModule(x_163); +lean_dec(x_163); x_165 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; x_166 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__14; lean_inc(x_110); @@ -7519,7 +7532,8 @@ x_24 = lean_ctor_get(x_22, 0); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); lean_dec(x_24); -x_26 = lean_environment_main_module(x_25); +x_26 = l_Lean_Environment_mainModule(x_25); +lean_dec(x_25); x_27 = l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__1; lean_inc(x_20); x_28 = lean_alloc_ctor(2, 2, 0); @@ -7584,7 +7598,8 @@ lean_dec(x_22); x_52 = lean_ctor_get(x_50, 0); lean_inc(x_52); lean_dec(x_50); -x_53 = lean_environment_main_module(x_52); +x_53 = l_Lean_Environment_mainModule(x_52); +lean_dec(x_52); x_54 = l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__1; lean_inc(x_20); x_55 = lean_alloc_ctor(2, 2, 0); @@ -8269,7 +8284,8 @@ x_21 = lean_ctor_get(x_18, 0); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec(x_21); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; x_25 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__14; lean_inc(x_16); @@ -8431,7 +8447,8 @@ lean_dec(x_18); x_94 = lean_ctor_get(x_92, 0); lean_inc(x_94); lean_dec(x_92); -x_95 = lean_environment_main_module(x_94); +x_95 = l_Lean_Environment_mainModule(x_94); +lean_dec(x_94); x_96 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; x_97 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__14; lean_inc(x_16); @@ -9069,8 +9086,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -9111,8 +9127,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -10068,7 +10083,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -10114,7 +10129,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -10167,6 +10182,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_Deriving_BEq_mkBEqInstanceHandler___spec__2(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -10372,6 +10388,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_Deriving_BEq_mkBEqInstanceHandler___spec__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/Basic.c b/stage0/stdlib/Lean/Elab/Deriving/Basic.c index c43b25505a07..491da4f38ea0 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Basic.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Basic.c @@ -172,7 +172,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_Term_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_Term_mkInst_x3f_go_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1906____closed__11; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeriving___boxed__const__1; @@ -1156,7 +1156,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -1338,7 +1339,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); diff --git a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c index b95f5b8d4bbd..8410b9355a65 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c @@ -86,7 +86,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__17; LEAN_EXPORT lean_object* l_Lean_isEnumType___at_Lean_Elab_Deriving_DecEq_mkDecEqInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__41; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__9; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_DecEq_mkDecEq___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -324,7 +324,7 @@ static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_DecEq_mkM static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__35; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_DecEq_mkEnumOfNat___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunctions___closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat_mkDecTree___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__17; @@ -749,7 +749,8 @@ x_20 = lean_ctor_get(x_18, 0); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); lean_dec(x_20); -x_22 = lean_environment_main_module(x_21); +x_22 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_23 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__4; lean_inc(x_16); x_24 = lean_alloc_ctor(2, 2, 0); @@ -822,7 +823,8 @@ lean_dec(x_18); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__4; lean_inc(x_16); x_57 = lean_alloc_ctor(2, 2, 0); @@ -1402,7 +1404,8 @@ x_22 = lean_ctor_get(x_19, 1); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__3; lean_inc(x_17); lean_ctor_set_tag(x_19, 2); @@ -1618,7 +1621,8 @@ lean_dec(x_19); x_121 = lean_ctor_get(x_119, 0); lean_inc(x_121); lean_dec(x_119); -x_122 = lean_environment_main_module(x_121); +x_122 = l_Lean_Environment_mainModule(x_121); +lean_dec(x_121); x_123 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__3; lean_inc(x_17); x_124 = lean_alloc_ctor(2, 2, 0); @@ -2044,7 +2048,8 @@ x_22 = lean_ctor_get(x_19, 1); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__1; lean_inc(x_17); lean_ctor_set_tag(x_19, 2); @@ -2179,7 +2184,8 @@ lean_dec(x_19); x_86 = lean_ctor_get(x_84, 0); lean_inc(x_86); lean_dec(x_84); -x_87 = lean_environment_main_module(x_86); +x_87 = l_Lean_Environment_mainModule(x_86); +lean_dec(x_86); x_88 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__1; lean_inc(x_17); x_89 = lean_alloc_ctor(2, 2, 0); @@ -2426,7 +2432,8 @@ x_15 = lean_ctor_get(x_13, 0); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); lean_dec(x_15); -x_17 = lean_environment_main_module(x_16); +x_17 = l_Lean_Environment_mainModule(x_16); +lean_dec(x_16); x_18 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___closed__3; lean_inc(x_12); lean_inc(x_17); @@ -2468,7 +2475,8 @@ lean_dec(x_13); x_34 = lean_ctor_get(x_32, 0); lean_inc(x_34); lean_dec(x_32); -x_35 = lean_environment_main_module(x_34); +x_35 = l_Lean_Environment_mainModule(x_34); +lean_dec(x_34); x_36 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___closed__3; lean_inc(x_12); lean_inc(x_35); @@ -5997,7 +6005,8 @@ x_72 = lean_ctor_get(x_69, 1); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; lean_inc(x_68); lean_inc(x_74); @@ -6176,7 +6185,8 @@ lean_dec(x_69); x_150 = lean_ctor_get(x_148, 0); lean_inc(x_150); lean_dec(x_148); -x_151 = lean_environment_main_module(x_150); +x_151 = l_Lean_Environment_mainModule(x_150); +lean_dec(x_150); x_152 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; lean_inc(x_68); lean_inc(x_151); @@ -6347,7 +6357,8 @@ if (lean_is_exclusive(x_216)) { x_220 = lean_ctor_get(x_217, 0); lean_inc(x_220); lean_dec(x_217); -x_221 = lean_environment_main_module(x_220); +x_221 = l_Lean_Environment_mainModule(x_220); +lean_dec(x_220); x_222 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; lean_inc(x_215); lean_inc(x_221); @@ -6556,7 +6567,8 @@ if (lean_is_exclusive(x_298)) { x_302 = lean_ctor_get(x_299, 0); lean_inc(x_302); lean_dec(x_299); -x_303 = lean_environment_main_module(x_302); +x_303 = l_Lean_Environment_mainModule(x_302); +lean_dec(x_302); x_304 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; lean_inc(x_297); lean_inc(x_303); @@ -6948,7 +6960,8 @@ if (lean_is_exclusive(x_411)) { x_415 = lean_ctor_get(x_412, 0); lean_inc(x_415); lean_dec(x_412); -x_416 = lean_environment_main_module(x_415); +x_416 = l_Lean_Environment_mainModule(x_415); +lean_dec(x_415); x_417 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; lean_inc(x_410); lean_inc(x_416); @@ -7381,7 +7394,8 @@ if (lean_is_exclusive(x_542)) { x_546 = lean_ctor_get(x_543, 0); lean_inc(x_546); lean_dec(x_543); -x_547 = lean_environment_main_module(x_546); +x_547 = l_Lean_Environment_mainModule(x_546); +lean_dec(x_546); x_548 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; lean_inc(x_541); lean_inc(x_547); @@ -8631,7 +8645,8 @@ x_21 = lean_ctor_get(x_18, 1); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___lambda__1___closed__2; x_25 = l_Lean_addMacroScope(x_23, x_24, x_17); x_26 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___lambda__1___closed__1; @@ -8828,7 +8843,8 @@ lean_dec(x_18); x_111 = lean_ctor_get(x_109, 0); lean_inc(x_111); lean_dec(x_109); -x_112 = lean_environment_main_module(x_111); +x_112 = l_Lean_Environment_mainModule(x_111); +lean_dec(x_111); x_113 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___lambda__1___closed__2; x_114 = l_Lean_addMacroScope(x_112, x_113, x_17); x_115 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___lambda__1___closed__1; @@ -10370,8 +10386,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -10412,8 +10427,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c index 664121b81ecf..833cdfa16c4d 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c +++ b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c @@ -99,7 +99,7 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___lamb static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__5; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__6; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__13; @@ -387,7 +387,7 @@ static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkF LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__8___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__10; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__9; @@ -874,7 +874,8 @@ x_20 = lean_ctor_get(x_18, 0); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); lean_dec(x_20); -x_22 = lean_environment_main_module(x_21); +x_22 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_23 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__7; lean_inc(x_16); x_24 = lean_alloc_ctor(2, 2, 0); @@ -968,7 +969,8 @@ lean_dec(x_18); x_57 = lean_ctor_get(x_55, 0); lean_inc(x_57); lean_dec(x_55); -x_58 = lean_environment_main_module(x_57); +x_58 = l_Lean_Environment_mainModule(x_57); +lean_dec(x_57); x_59 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__7; lean_inc(x_16); x_60 = lean_alloc_ctor(2, 2, 0); @@ -1637,7 +1639,8 @@ x_33 = lean_ctor_get(x_30, 1); x_34 = lean_ctor_get(x_32, 0); lean_inc(x_34); lean_dec(x_32); -x_35 = lean_environment_main_module(x_34); +x_35 = l_Lean_Environment_mainModule(x_34); +lean_dec(x_34); x_36 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__3; lean_inc(x_28); lean_ctor_set_tag(x_30, 2); @@ -1728,7 +1731,8 @@ lean_dec(x_30); x_74 = lean_ctor_get(x_72, 0); lean_inc(x_74); lean_dec(x_72); -x_75 = lean_environment_main_module(x_74); +x_75 = l_Lean_Environment_mainModule(x_74); +lean_dec(x_74); x_76 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__3; lean_inc(x_28); x_77 = lean_alloc_ctor(2, 2, 0); @@ -1830,7 +1834,8 @@ x_120 = lean_ctor_get(x_117, 1); x_121 = lean_ctor_get(x_119, 0); lean_inc(x_121); lean_dec(x_119); -x_122 = lean_environment_main_module(x_121); +x_122 = l_Lean_Environment_mainModule(x_121); +lean_dec(x_121); x_123 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__23; x_124 = l_Lean_addMacroScope(x_122, x_123, x_116); x_125 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__22; @@ -1874,7 +1879,8 @@ lean_dec(x_117); x_142 = lean_ctor_get(x_140, 0); lean_inc(x_142); lean_dec(x_140); -x_143 = lean_environment_main_module(x_142); +x_143 = l_Lean_Environment_mainModule(x_142); +lean_dec(x_142); x_144 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__23; x_145 = l_Lean_addMacroScope(x_143, x_144, x_116); x_146 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__22; @@ -2182,7 +2188,8 @@ x_27 = lean_ctor_get(x_25, 0); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); lean_dec(x_27); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; lean_inc(x_24); lean_inc(x_29); @@ -2256,7 +2263,8 @@ lean_dec(x_25); x_60 = lean_ctor_get(x_58, 0); lean_inc(x_60); lean_dec(x_58); -x_61 = lean_environment_main_module(x_60); +x_61 = l_Lean_Environment_mainModule(x_60); +lean_dec(x_60); x_62 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; lean_inc(x_24); lean_inc(x_61); @@ -2396,7 +2404,8 @@ if (lean_is_exclusive(x_108)) { x_112 = lean_ctor_get(x_109, 0); lean_inc(x_112); lean_dec(x_109); -x_113 = lean_environment_main_module(x_112); +x_113 = l_Lean_Environment_mainModule(x_112); +lean_dec(x_112); x_114 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; lean_inc(x_107); lean_inc(x_113); @@ -4631,7 +4640,8 @@ lean_dec(x_27); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__11; x_33 = l_Lean_addMacroScope(x_31, x_32, x_26); x_34 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__10; @@ -4906,7 +4916,8 @@ lean_dec(x_108); x_111 = lean_ctor_get(x_109, 0); lean_inc(x_111); lean_dec(x_109); -x_112 = lean_environment_main_module(x_111); +x_112 = l_Lean_Environment_mainModule(x_111); +lean_dec(x_111); x_113 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__11; x_114 = l_Lean_addMacroScope(x_112, x_113, x_107); x_115 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__10; @@ -5186,7 +5197,8 @@ lean_dec(x_108); x_111 = lean_ctor_get(x_109, 0); lean_inc(x_111); lean_dec(x_109); -x_112 = lean_environment_main_module(x_111); +x_112 = l_Lean_Environment_mainModule(x_111); +lean_dec(x_111); x_113 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__11; x_114 = l_Lean_addMacroScope(x_112, x_113, x_107); x_115 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__10; @@ -5440,7 +5452,8 @@ x_34 = lean_ctor_get(x_32, 0); x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); lean_dec(x_34); -x_36 = lean_environment_main_module(x_35); +x_36 = l_Lean_Environment_mainModule(x_35); +lean_dec(x_35); x_37 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; lean_inc(x_31); lean_inc(x_36); @@ -5548,7 +5561,8 @@ lean_dec(x_32); x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); lean_dec(x_82); -x_85 = lean_environment_main_module(x_84); +x_85 = l_Lean_Environment_mainModule(x_84); +lean_dec(x_84); x_86 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; lean_inc(x_31); lean_inc(x_85); @@ -5680,7 +5694,8 @@ if (lean_is_exclusive(x_138)) { x_142 = lean_ctor_get(x_139, 0); lean_inc(x_142); lean_dec(x_139); -x_143 = lean_environment_main_module(x_142); +x_143 = l_Lean_Environment_mainModule(x_142); +lean_dec(x_142); x_144 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; lean_inc(x_137); lean_inc(x_143); @@ -5815,7 +5830,8 @@ x_203 = lean_ctor_get(x_201, 0); x_204 = lean_ctor_get(x_203, 0); lean_inc(x_204); lean_dec(x_203); -x_205 = lean_environment_main_module(x_204); +x_205 = l_Lean_Environment_mainModule(x_204); +lean_dec(x_204); x_206 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; x_207 = l_Lean_addMacroScope(x_205, x_206, x_200); x_208 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__4; @@ -5908,7 +5924,8 @@ lean_dec(x_201); x_246 = lean_ctor_get(x_244, 0); lean_inc(x_246); lean_dec(x_244); -x_247 = lean_environment_main_module(x_246); +x_247 = l_Lean_Environment_mainModule(x_246); +lean_dec(x_246); x_248 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; x_249 = l_Lean_addMacroScope(x_247, x_248, x_200); x_250 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__4; @@ -6025,7 +6042,8 @@ if (lean_is_exclusive(x_293)) { x_297 = lean_ctor_get(x_294, 0); lean_inc(x_297); lean_dec(x_294); -x_298 = lean_environment_main_module(x_297); +x_298 = l_Lean_Environment_mainModule(x_297); +lean_dec(x_297); x_299 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; x_300 = l_Lean_addMacroScope(x_298, x_299, x_292); x_301 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__4; @@ -6177,7 +6195,8 @@ lean_dec(x_424); x_427 = lean_ctor_get(x_425, 0); lean_inc(x_427); lean_dec(x_425); -x_428 = lean_environment_main_module(x_427); +x_428 = l_Lean_Environment_mainModule(x_427); +lean_dec(x_427); x_429 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__11; x_430 = l_Lean_addMacroScope(x_428, x_429, x_423); x_431 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__10; @@ -6238,7 +6257,8 @@ x_355 = lean_ctor_get(x_353, 0); x_356 = lean_ctor_get(x_355, 0); lean_inc(x_356); lean_dec(x_355); -x_357 = lean_environment_main_module(x_356); +x_357 = l_Lean_Environment_mainModule(x_356); +lean_dec(x_356); x_358 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; x_359 = l_Lean_addMacroScope(x_357, x_358, x_352); x_360 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__4; @@ -6313,7 +6333,8 @@ lean_dec(x_353); x_388 = lean_ctor_get(x_386, 0); lean_inc(x_388); lean_dec(x_386); -x_389 = lean_environment_main_module(x_388); +x_389 = l_Lean_Environment_mainModule(x_388); +lean_dec(x_388); x_390 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; x_391 = l_Lean_addMacroScope(x_389, x_390, x_352); x_392 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__4; @@ -6419,7 +6440,8 @@ x_466 = lean_ctor_get(x_464, 0); x_467 = lean_ctor_get(x_466, 0); lean_inc(x_467); lean_dec(x_466); -x_468 = lean_environment_main_module(x_467); +x_468 = l_Lean_Environment_mainModule(x_467); +lean_dec(x_467); x_469 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; x_470 = l_Lean_addMacroScope(x_468, x_469, x_463); x_471 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__4; @@ -6512,7 +6534,8 @@ lean_dec(x_464); x_509 = lean_ctor_get(x_507, 0); lean_inc(x_509); lean_dec(x_507); -x_510 = lean_environment_main_module(x_509); +x_510 = l_Lean_Environment_mainModule(x_509); +lean_dec(x_509); x_511 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; x_512 = l_Lean_addMacroScope(x_510, x_511, x_463); x_513 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__4; @@ -6629,7 +6652,8 @@ if (lean_is_exclusive(x_556)) { x_560 = lean_ctor_get(x_557, 0); lean_inc(x_560); lean_dec(x_557); -x_561 = lean_environment_main_module(x_560); +x_561 = l_Lean_Environment_mainModule(x_560); +lean_dec(x_560); x_562 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__5; x_563 = l_Lean_addMacroScope(x_561, x_562, x_555); x_564 = l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__4; @@ -6743,7 +6767,8 @@ x_608 = lean_ctor_get(x_606, 0); x_609 = lean_ctor_get(x_608, 0); lean_inc(x_609); lean_dec(x_608); -x_610 = lean_environment_main_module(x_609); +x_610 = l_Lean_Environment_mainModule(x_609); +lean_dec(x_609); x_611 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__11; x_612 = l_Lean_addMacroScope(x_610, x_611, x_605); x_613 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__10; @@ -6776,7 +6801,8 @@ lean_dec(x_606); x_624 = lean_ctor_get(x_622, 0); lean_inc(x_624); lean_dec(x_622); -x_625 = lean_environment_main_module(x_624); +x_625 = l_Lean_Environment_mainModule(x_624); +lean_dec(x_624); x_626 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__11; x_627 = l_Lean_addMacroScope(x_625, x_626, x_605); x_628 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__10; @@ -7718,7 +7744,8 @@ x_27 = lean_ctor_get(x_24, 1); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__3; lean_inc(x_23); lean_inc(x_29); @@ -7770,7 +7797,8 @@ x_52 = lean_ctor_get(x_49, 1); x_53 = lean_ctor_get(x_51, 0); lean_inc(x_53); lean_dec(x_51); -x_54 = lean_environment_main_module(x_53); +x_54 = l_Lean_Environment_mainModule(x_53); +lean_dec(x_53); x_55 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__14; lean_inc(x_23); lean_inc(x_54); @@ -8034,7 +8062,8 @@ lean_dec(x_49); x_168 = lean_ctor_get(x_166, 0); lean_inc(x_168); lean_dec(x_166); -x_169 = lean_environment_main_module(x_168); +x_169 = l_Lean_Environment_mainModule(x_168); +lean_dec(x_168); x_170 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__14; lean_inc(x_23); lean_inc(x_169); @@ -8299,7 +8328,8 @@ lean_dec(x_24); x_284 = lean_ctor_get(x_282, 0); lean_inc(x_284); lean_dec(x_282); -x_285 = lean_environment_main_module(x_284); +x_285 = l_Lean_Environment_mainModule(x_284); +lean_dec(x_284); x_286 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__3; lean_inc(x_23); lean_inc(x_285); @@ -8357,7 +8387,8 @@ if (lean_is_exclusive(x_306)) { x_310 = lean_ctor_get(x_307, 0); lean_inc(x_310); lean_dec(x_307); -x_311 = lean_environment_main_module(x_310); +x_311 = l_Lean_Environment_mainModule(x_310); +lean_dec(x_310); x_312 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__14; lean_inc(x_23); lean_inc(x_311); @@ -9982,7 +10013,8 @@ x_34 = lean_ctor_get(x_31, 1); x_35 = lean_ctor_get(x_33, 0); lean_inc(x_35); lean_dec(x_33); -x_36 = lean_environment_main_module(x_35); +x_36 = l_Lean_Environment_mainModule(x_35); +lean_dec(x_35); x_37 = l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__3; lean_inc(x_30); lean_inc(x_36); @@ -10064,7 +10096,8 @@ lean_dec(x_31); x_70 = lean_ctor_get(x_68, 0); lean_inc(x_70); lean_dec(x_68); -x_71 = lean_environment_main_module(x_70); +x_71 = l_Lean_Environment_mainModule(x_70); +lean_dec(x_70); x_72 = l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__3; lean_inc(x_30); lean_inc(x_71); @@ -10156,7 +10189,8 @@ x_111 = lean_ctor_get(x_108, 1); x_112 = lean_ctor_get(x_110, 0); lean_inc(x_112); lean_dec(x_110); -x_113 = lean_environment_main_module(x_112); +x_113 = l_Lean_Environment_mainModule(x_112); +lean_dec(x_112); x_114 = l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__11; x_115 = l_Lean_addMacroScope(x_113, x_114, x_107); x_116 = lean_box(0); @@ -10227,7 +10261,8 @@ lean_dec(x_108); x_142 = lean_ctor_get(x_140, 0); lean_inc(x_142); lean_dec(x_140); -x_143 = lean_environment_main_module(x_142); +x_143 = l_Lean_Environment_mainModule(x_142); +lean_dec(x_142); x_144 = l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__11; x_145 = l_Lean_addMacroScope(x_143, x_144, x_107); x_146 = lean_box(0); @@ -10325,7 +10360,8 @@ if (lean_is_exclusive(x_179)) { x_183 = lean_ctor_get(x_180, 0); lean_inc(x_183); lean_dec(x_180); -x_184 = lean_environment_main_module(x_183); +x_184 = l_Lean_Environment_mainModule(x_183); +lean_dec(x_183); x_185 = l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__3; lean_inc(x_178); lean_inc(x_184); @@ -10427,7 +10463,8 @@ if (lean_is_exclusive(x_222)) { x_226 = lean_ctor_get(x_223, 0); lean_inc(x_226); lean_dec(x_223); -x_227 = lean_environment_main_module(x_226); +x_227 = l_Lean_Environment_mainModule(x_226); +lean_dec(x_226); x_228 = l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__11; x_229 = l_Lean_addMacroScope(x_227, x_228, x_221); x_230 = lean_box(0); @@ -10837,7 +10874,8 @@ x_21 = lean_ctor_get(x_19, 0); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec(x_21); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__2___closed__2; lean_inc(x_1); x_25 = lean_alloc_ctor(1, 2, 0); @@ -11048,7 +11086,8 @@ lean_dec(x_19); x_114 = lean_ctor_get(x_112, 0); lean_inc(x_114); lean_dec(x_112); -x_115 = lean_environment_main_module(x_114); +x_115 = l_Lean_Environment_mainModule(x_114); +lean_dec(x_114); x_116 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__2___closed__2; lean_inc(x_1); x_117 = lean_alloc_ctor(1, 2, 0); @@ -11501,7 +11540,8 @@ lean_dec(x_45); x_48 = lean_ctor_get(x_46, 0); lean_inc(x_48); lean_dec(x_46); -x_49 = lean_environment_main_module(x_48); +x_49 = l_Lean_Environment_mainModule(x_48); +lean_dec(x_48); x_50 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__4; x_51 = l_Lean_addMacroScope(x_49, x_50, x_44); x_52 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__3; @@ -11536,7 +11576,8 @@ x_63 = lean_ctor_get(x_60, 1); x_64 = lean_ctor_get(x_62, 0); lean_inc(x_64); lean_dec(x_62); -x_65 = lean_environment_main_module(x_64); +x_65 = l_Lean_Environment_mainModule(x_64); +lean_dec(x_64); x_66 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__12; x_67 = l_Lean_addMacroScope(x_65, x_66, x_59); x_68 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__11; @@ -11597,7 +11638,8 @@ lean_dec(x_60); x_92 = lean_ctor_get(x_90, 0); lean_inc(x_92); lean_dec(x_90); -x_93 = lean_environment_main_module(x_92); +x_93 = l_Lean_Environment_mainModule(x_92); +lean_dec(x_92); x_94 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__12; x_95 = l_Lean_addMacroScope(x_93, x_94, x_59); x_96 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__11; @@ -11681,7 +11723,8 @@ lean_dec(x_127); x_130 = lean_ctor_get(x_128, 0); lean_inc(x_130); lean_dec(x_128); -x_131 = lean_environment_main_module(x_130); +x_131 = l_Lean_Environment_mainModule(x_130); +lean_dec(x_130); x_132 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__4; x_133 = l_Lean_addMacroScope(x_131, x_132, x_126); x_134 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__3; @@ -11722,7 +11765,8 @@ if (lean_is_exclusive(x_142)) { x_146 = lean_ctor_get(x_143, 0); lean_inc(x_146); lean_dec(x_143); -x_147 = lean_environment_main_module(x_146); +x_147 = l_Lean_Environment_mainModule(x_146); +lean_dec(x_146); x_148 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__12; x_149 = l_Lean_addMacroScope(x_147, x_148, x_141); x_150 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__11; @@ -12540,7 +12584,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_29); +x_30 = l_Lean_Environment_mainModule(x_29); +lean_dec(x_29); x_31 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___spec__1___closed__4; x_32 = l_Lean_addMacroScope(x_30, x_31, x_24); lean_inc(x_3); @@ -12626,7 +12671,8 @@ lean_dec(x_25); x_62 = lean_ctor_get(x_60, 0); lean_inc(x_62); lean_dec(x_60); -x_63 = lean_environment_main_module(x_62); +x_63 = l_Lean_Environment_mainModule(x_62); +lean_dec(x_62); x_64 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___spec__1___closed__4; x_65 = l_Lean_addMacroScope(x_63, x_64, x_24); lean_inc(x_3); @@ -12745,7 +12791,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_29); +x_30 = l_Lean_Environment_mainModule(x_29); +lean_dec(x_29); x_31 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___spec__1___closed__4; x_32 = l_Lean_addMacroScope(x_30, x_31, x_24); lean_inc(x_3); @@ -12831,7 +12878,8 @@ lean_dec(x_25); x_62 = lean_ctor_get(x_60, 0); lean_inc(x_62); lean_dec(x_60); -x_63 = lean_environment_main_module(x_62); +x_63 = l_Lean_Environment_mainModule(x_62); +lean_dec(x_62); x_64 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___spec__1___closed__4; x_65 = l_Lean_addMacroScope(x_63, x_64, x_24); lean_inc(x_3); @@ -13069,7 +13117,8 @@ x_23 = lean_ctor_get(x_20, 1); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = lean_box(0); x_27 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___closed__5; x_28 = l_Lean_addMacroScope(x_25, x_27, x_19); @@ -13216,7 +13265,8 @@ lean_dec(x_20); x_67 = lean_ctor_get(x_65, 0); lean_inc(x_67); lean_dec(x_65); -x_68 = lean_environment_main_module(x_67); +x_68 = l_Lean_Environment_mainModule(x_67); +lean_dec(x_67); x_69 = lean_box(0); x_70 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___closed__5; x_71 = l_Lean_addMacroScope(x_68, x_70, x_19); @@ -13813,7 +13863,8 @@ x_29 = lean_ctor_get(x_27, 0); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); lean_dec(x_29); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_33 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_25); @@ -13906,7 +13957,8 @@ lean_dec(x_27); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_76 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_25); @@ -14048,7 +14100,8 @@ x_132 = lean_ctor_get(x_130, 0); x_133 = lean_ctor_get(x_132, 0); lean_inc(x_133); lean_dec(x_132); -x_134 = lean_environment_main_module(x_133); +x_134 = l_Lean_Environment_mainModule(x_133); +lean_dec(x_133); x_135 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_136 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_128); @@ -14151,7 +14204,8 @@ lean_dec(x_130); x_180 = lean_ctor_get(x_178, 0); lean_inc(x_180); lean_dec(x_178); -x_181 = lean_environment_main_module(x_180); +x_181 = l_Lean_Environment_mainModule(x_180); +lean_dec(x_180); x_182 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_183 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_128); @@ -14279,7 +14333,8 @@ if (lean_is_exclusive(x_232)) { x_236 = lean_ctor_get(x_233, 0); lean_inc(x_236); lean_dec(x_233); -x_237 = lean_environment_main_module(x_236); +x_237 = l_Lean_Environment_mainModule(x_236); +lean_dec(x_236); x_238 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_239 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_230); @@ -14874,7 +14929,8 @@ x_302 = lean_ctor_get(x_300, 0); x_303 = lean_ctor_get(x_302, 0); lean_inc(x_303); lean_dec(x_302); -x_304 = lean_environment_main_module(x_303); +x_304 = l_Lean_Environment_mainModule(x_303); +lean_dec(x_303); x_305 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_306 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_298); @@ -14984,7 +15040,8 @@ lean_dec(x_300); x_353 = lean_ctor_get(x_351, 0); lean_inc(x_353); lean_dec(x_351); -x_354 = lean_environment_main_module(x_353); +x_354 = l_Lean_Environment_mainModule(x_353); +lean_dec(x_353); x_355 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_356 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_298); @@ -15119,7 +15176,8 @@ if (lean_is_exclusive(x_408)) { x_412 = lean_ctor_get(x_409, 0); lean_inc(x_412); lean_dec(x_409); -x_413 = lean_environment_main_module(x_412); +x_413 = l_Lean_Environment_mainModule(x_412); +lean_dec(x_412); x_414 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_415 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_406); @@ -15302,7 +15360,8 @@ x_43 = lean_ctor_get(x_41, 0); x_44 = lean_ctor_get(x_43, 0); lean_inc(x_44); lean_dec(x_43); -x_45 = lean_environment_main_module(x_44); +x_45 = l_Lean_Environment_mainModule(x_44); +lean_dec(x_44); x_46 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_47 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_39); @@ -15422,7 +15481,8 @@ lean_dec(x_41); x_98 = lean_ctor_get(x_96, 0); lean_inc(x_98); lean_dec(x_96); -x_99 = lean_environment_main_module(x_98); +x_99 = l_Lean_Environment_mainModule(x_98); +lean_dec(x_98); x_100 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_101 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_39); @@ -15567,7 +15627,8 @@ if (lean_is_exclusive(x_157)) { x_161 = lean_ctor_get(x_158, 0); lean_inc(x_161); lean_dec(x_158); -x_162 = lean_environment_main_module(x_161); +x_162 = l_Lean_Environment_mainModule(x_161); +lean_dec(x_161); x_163 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_164 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_155); @@ -15733,7 +15794,8 @@ if (lean_is_exclusive(x_225)) { x_229 = lean_ctor_get(x_226, 0); lean_inc(x_229); lean_dec(x_226); -x_230 = lean_environment_main_module(x_229); +x_230 = l_Lean_Environment_mainModule(x_229); +lean_dec(x_229); x_231 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; x_232 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_223); @@ -17487,7 +17549,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -17533,7 +17595,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -17586,6 +17648,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__3(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -17822,6 +17885,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__3(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c index 16f739a4c26c..50247302adba 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c @@ -58,7 +58,7 @@ static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___cl static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__16; static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__1___closed__2; static lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2340____closed__3; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2340____closed__9; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -165,7 +165,7 @@ lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__18; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__20; static lean_object* l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -1104,7 +1104,8 @@ x_64 = lean_ctor_get(x_62, 0); x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); lean_dec(x_64); -x_66 = lean_environment_main_module(x_65); +x_66 = l_Lean_Environment_mainModule(x_65); +lean_dec(x_65); x_67 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; lean_inc(x_61); lean_inc(x_66); @@ -1169,7 +1170,8 @@ lean_dec(x_62); x_93 = lean_ctor_get(x_91, 0); lean_inc(x_93); lean_dec(x_91); -x_94 = lean_environment_main_module(x_93); +x_94 = l_Lean_Environment_mainModule(x_93); +lean_dec(x_93); x_95 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; lean_inc(x_61); lean_inc(x_94); @@ -1250,7 +1252,8 @@ x_128 = lean_ctor_get(x_126, 0); x_129 = lean_ctor_get(x_128, 0); lean_inc(x_129); lean_dec(x_128); -x_130 = lean_environment_main_module(x_129); +x_130 = l_Lean_Environment_mainModule(x_129); +lean_dec(x_129); x_131 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; x_132 = l_Lean_addMacroScope(x_130, x_131, x_125); x_133 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__4; @@ -1337,7 +1340,8 @@ lean_dec(x_126); x_166 = lean_ctor_get(x_164, 0); lean_inc(x_166); lean_dec(x_164); -x_167 = lean_environment_main_module(x_166); +x_167 = l_Lean_Environment_mainModule(x_166); +lean_dec(x_166); x_168 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; x_169 = l_Lean_addMacroScope(x_167, x_168, x_125); x_170 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__4; @@ -1447,7 +1451,8 @@ if (lean_is_exclusive(x_208)) { x_212 = lean_ctor_get(x_209, 0); lean_inc(x_212); lean_dec(x_209); -x_213 = lean_environment_main_module(x_212); +x_213 = l_Lean_Environment_mainModule(x_212); +lean_dec(x_212); x_214 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; x_215 = l_Lean_addMacroScope(x_213, x_214, x_207); x_216 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__4; @@ -1559,7 +1564,8 @@ x_257 = lean_ctor_get(x_255, 0); x_258 = lean_ctor_get(x_257, 0); lean_inc(x_258); lean_dec(x_257); -x_259 = lean_environment_main_module(x_258); +x_259 = l_Lean_Environment_mainModule(x_258); +lean_dec(x_258); x_260 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; lean_inc(x_254); lean_inc(x_259); @@ -1624,7 +1630,8 @@ lean_dec(x_255); x_286 = lean_ctor_get(x_284, 0); lean_inc(x_286); lean_dec(x_284); -x_287 = lean_environment_main_module(x_286); +x_287 = l_Lean_Environment_mainModule(x_286); +lean_dec(x_286); x_288 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__5; lean_inc(x_254); lean_inc(x_287); @@ -3867,7 +3874,8 @@ x_21 = lean_ctor_get(x_19, 0); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec(x_21); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__9; x_25 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__6___lambda__1___closed__8; lean_inc(x_17); @@ -3960,7 +3968,8 @@ lean_dec(x_19); x_65 = lean_ctor_get(x_63, 0); lean_inc(x_65); lean_dec(x_63); -x_66 = lean_environment_main_module(x_65); +x_66 = l_Lean_Environment_mainModule(x_65); +lean_dec(x_65); x_67 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__9; x_68 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__6___lambda__1___closed__8; lean_inc(x_17); @@ -4066,7 +4075,8 @@ x_114 = lean_ctor_get(x_112, 0); x_115 = lean_ctor_get(x_114, 0); lean_inc(x_115); lean_dec(x_114); -x_116 = lean_environment_main_module(x_115); +x_116 = l_Lean_Environment_mainModule(x_115); +lean_dec(x_115); x_117 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__9; x_118 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__6___lambda__1___closed__8; lean_inc(x_110); @@ -4169,7 +4179,8 @@ lean_dec(x_112); x_163 = lean_ctor_get(x_161, 0); lean_inc(x_163); lean_dec(x_161); -x_164 = lean_environment_main_module(x_163); +x_164 = l_Lean_Environment_mainModule(x_163); +lean_dec(x_163); x_165 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___closed__9; x_166 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__6___lambda__1___closed__8; lean_inc(x_110); @@ -5385,7 +5396,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -5431,7 +5442,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -5484,6 +5495,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_Deriving_Hashable_mkHashableHandler___spec__3(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -5702,6 +5714,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_Deriving_Hashable_mkHashableHandler___spec__3(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c index f6340aa909f4..9cb630ba2b0a 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c @@ -73,7 +73,7 @@ lean_object* lean_array_fget(lean_object*, lean_object*); extern lean_object* l_Lean_ForEachExprWhere_initCache; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -234,7 +234,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_m static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__6; static lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__1___closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_collectUsedLocalsInsts___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -6790,7 +6790,8 @@ x_52 = lean_ctor_get(x_49, 1); x_53 = lean_ctor_get(x_51, 0); lean_inc(x_53); lean_dec(x_51); -x_54 = lean_environment_main_module(x_53); +x_54 = l_Lean_Environment_mainModule(x_53); +lean_dec(x_53); x_55 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; lean_inc(x_29); lean_ctor_set_tag(x_49, 2); @@ -6841,7 +6842,8 @@ lean_dec(x_49); x_74 = lean_ctor_get(x_72, 0); lean_inc(x_74); lean_dec(x_72); -x_75 = lean_environment_main_module(x_74); +x_75 = l_Lean_Environment_mainModule(x_74); +lean_dec(x_74); x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; lean_inc(x_29); x_77 = lean_alloc_ctor(2, 2, 0); @@ -6956,7 +6958,8 @@ if (lean_is_exclusive(x_111)) { x_115 = lean_ctor_get(x_112, 0); lean_inc(x_115); lean_dec(x_112); -x_116 = lean_environment_main_module(x_115); +x_116 = l_Lean_Environment_mainModule(x_115); +lean_dec(x_115); x_117 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; lean_inc(x_29); if (lean_is_scalar(x_114)) { @@ -7102,7 +7105,8 @@ if (lean_is_exclusive(x_162)) { x_166 = lean_ctor_get(x_163, 0); lean_inc(x_166); lean_dec(x_163); -x_167 = lean_environment_main_module(x_166); +x_167 = l_Lean_Environment_mainModule(x_166); +lean_dec(x_166); x_168 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; lean_inc(x_141); if (lean_is_scalar(x_165)) { @@ -7326,7 +7330,8 @@ x_27 = lean_ctor_get(x_24, 1); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__4; x_31 = l_Lean_addMacroScope(x_29, x_30, x_23); lean_inc(x_2); @@ -7365,7 +7370,8 @@ lean_dec(x_24); x_41 = lean_ctor_get(x_39, 0); lean_inc(x_41); lean_dec(x_39); -x_42 = lean_environment_main_module(x_41); +x_42 = l_Lean_Environment_mainModule(x_41); +lean_dec(x_41); x_43 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__4; x_44 = l_Lean_addMacroScope(x_42, x_43, x_23); lean_inc(x_2); @@ -7892,7 +7898,8 @@ x_39 = lean_ctor_get(x_36, 1); x_40 = lean_ctor_get(x_38, 0); lean_inc(x_40); lean_dec(x_38); -x_41 = lean_environment_main_module(x_40); +x_41 = l_Lean_Environment_mainModule(x_40); +lean_dec(x_40); x_42 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; x_43 = l_Lean_addMacroScope(x_41, x_42, x_35); x_44 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18; @@ -8562,7 +8569,8 @@ lean_dec(x_36); x_303 = lean_ctor_get(x_301, 0); lean_inc(x_303); lean_dec(x_301); -x_304 = lean_environment_main_module(x_303); +x_304 = l_Lean_Environment_mainModule(x_303); +lean_dec(x_303); x_305 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; x_306 = l_Lean_addMacroScope(x_304, x_305, x_35); x_307 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18; @@ -8821,7 +8829,8 @@ if (lean_is_exclusive(x_396)) { x_400 = lean_ctor_get(x_397, 0); lean_inc(x_400); lean_dec(x_397); -x_401 = lean_environment_main_module(x_400); +x_401 = l_Lean_Environment_mainModule(x_400); +lean_dec(x_400); x_402 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; x_403 = l_Lean_addMacroScope(x_401, x_402, x_395); x_404 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18; @@ -9096,7 +9105,8 @@ if (lean_is_exclusive(x_496)) { x_500 = lean_ctor_get(x_497, 0); lean_inc(x_500); lean_dec(x_497); -x_501 = lean_environment_main_module(x_500); +x_501 = l_Lean_Environment_mainModule(x_500); +lean_dec(x_500); x_502 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; x_503 = l_Lean_addMacroScope(x_501, x_502, x_495); x_504 = l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18; @@ -11860,8 +11870,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -11902,8 +11911,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -12891,7 +12899,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -12937,7 +12945,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -12990,6 +12998,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_mkInhabitedInstanceHandler___spec__2(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -13229,6 +13238,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_mkInhabitedInstanceHandler___spec__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c b/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c index 7795a0790c2d..711360202f25 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c @@ -43,7 +43,7 @@ lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_ob static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__3; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___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_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__20; static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -135,7 +135,7 @@ static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Nonempty___hyg_ static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__26; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__3___closed__21; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__39; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__3___closed__9; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_mkNonemptyInstanceHandler___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); @@ -780,7 +780,8 @@ x_72 = lean_ctor_get(x_69, 1); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__3___closed__13; lean_inc(x_39); lean_ctor_set_tag(x_69, 2); @@ -826,7 +827,8 @@ lean_dec(x_69); x_91 = lean_ctor_get(x_89, 0); lean_inc(x_91); lean_dec(x_89); -x_92 = lean_environment_main_module(x_91); +x_92 = l_Lean_Environment_mainModule(x_91); +lean_dec(x_91); x_93 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__3___closed__13; lean_inc(x_39); x_94 = lean_alloc_ctor(2, 2, 0); @@ -887,7 +889,8 @@ if (lean_is_exclusive(x_110)) { x_114 = lean_ctor_get(x_111, 0); lean_inc(x_114); lean_dec(x_111); -x_115 = lean_environment_main_module(x_114); +x_115 = l_Lean_Environment_mainModule(x_114); +lean_dec(x_114); x_116 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__3___closed__13; lean_inc(x_39); if (lean_is_scalar(x_113)) { @@ -1164,7 +1167,8 @@ if (lean_is_exclusive(x_173)) { x_177 = lean_ctor_get(x_174, 0); lean_inc(x_177); lean_dec(x_174); -x_178 = lean_environment_main_module(x_177); +x_178 = l_Lean_Environment_mainModule(x_177); +lean_dec(x_177); x_179 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__3___closed__13; lean_inc(x_39); if (lean_is_scalar(x_176)) { @@ -1477,7 +1481,8 @@ if (lean_is_exclusive(x_246)) { x_250 = lean_ctor_get(x_247, 0); lean_inc(x_250); lean_dec(x_247); -x_251 = lean_environment_main_module(x_250); +x_251 = l_Lean_Environment_mainModule(x_250); +lean_dec(x_250); x_252 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__3___closed__13; lean_inc(x_216); if (lean_is_scalar(x_249)) { @@ -1921,7 +1926,8 @@ x_23 = lean_ctor_get(x_20, 1); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__4___closed__4; lean_inc(x_18); lean_ctor_set_tag(x_20, 2); @@ -1984,7 +1990,8 @@ lean_dec(x_20); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__4___closed__4; lean_inc(x_18); x_57 = lean_alloc_ctor(2, 2, 0); @@ -2787,7 +2794,8 @@ x_35 = lean_ctor_get(x_33, 0); x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); lean_dec(x_35); -x_37 = lean_environment_main_module(x_36); +x_37 = l_Lean_Environment_mainModule(x_36); +lean_dec(x_36); x_38 = l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__6; lean_inc(x_31); lean_ctor_set_tag(x_25, 2); @@ -2978,7 +2986,8 @@ lean_dec(x_33); x_126 = lean_ctor_get(x_124, 0); lean_inc(x_126); lean_dec(x_124); -x_127 = lean_environment_main_module(x_126); +x_127 = l_Lean_Environment_mainModule(x_126); +lean_dec(x_126); x_128 = l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__6; lean_inc(x_31); lean_ctor_set_tag(x_25, 2); @@ -3194,7 +3203,8 @@ if (lean_is_exclusive(x_221)) { x_225 = lean_ctor_get(x_222, 0); lean_inc(x_225); lean_dec(x_222); -x_226 = lean_environment_main_module(x_225); +x_226 = l_Lean_Environment_mainModule(x_225); +lean_dec(x_225); x_227 = l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__6; lean_inc(x_219); x_228 = lean_alloc_ctor(2, 2, 0); @@ -3435,7 +3445,8 @@ if (lean_is_exclusive(x_328)) { x_332 = lean_ctor_get(x_329, 0); lean_inc(x_332); lean_dec(x_329); -x_333 = lean_environment_main_module(x_332); +x_333 = l_Lean_Environment_mainModule(x_332); +lean_dec(x_332); x_334 = l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__6; lean_inc(x_326); if (lean_is_scalar(x_323)) { @@ -3948,7 +3959,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -3994,7 +4005,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -4047,6 +4058,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_mkNonemptyInstanceHandler___spec__2(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -4252,6 +4264,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_mkNonemptyInstanceHandler___spec__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/Ord.c b/stage0/stdlib/Lean/Elab/Deriving/Ord.c index ec7ae4ef9ad5..d4da03770be5 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Ord.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Ord.c @@ -63,7 +63,7 @@ static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed_ static lean_object* l_Lean_Elab_Deriving_Ord_mkOrdHeader___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13; static lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2789____closed__7; static lean_object* l_Lean_Elab_Deriving_Ord_mkMatch___closed__2; @@ -198,7 +198,7 @@ lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__3; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__1; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__25; @@ -1021,7 +1021,8 @@ x_19 = lean_ctor_get(x_16, 1); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__7; lean_inc(x_15); lean_inc(x_21); @@ -1080,7 +1081,8 @@ lean_dec(x_16); x_46 = lean_ctor_get(x_44, 0); lean_inc(x_46); lean_dec(x_44); -x_47 = lean_environment_main_module(x_46); +x_47 = l_Lean_Environment_mainModule(x_46); +lean_dec(x_46); x_48 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__7; lean_inc(x_15); lean_inc(x_47); @@ -2235,7 +2237,8 @@ x_96 = lean_ctor_get(x_93, 1); x_97 = lean_ctor_get(x_95, 0); lean_inc(x_97); lean_dec(x_95); -x_98 = lean_environment_main_module(x_97); +x_98 = l_Lean_Environment_mainModule(x_97); +lean_dec(x_97); x_99 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_92); x_100 = l_Lean_addMacroScope(x_98, x_99, x_92); @@ -2305,7 +2308,8 @@ x_126 = lean_ctor_get(x_123, 1); x_127 = lean_ctor_get(x_125, 0); lean_inc(x_127); lean_dec(x_125); -x_128 = lean_environment_main_module(x_127); +x_128 = l_Lean_Environment_mainModule(x_127); +lean_dec(x_127); x_129 = lean_array_size(x_85); x_130 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_129, x_113, x_85); x_131 = l_Lean_mkSepArray(x_130, x_115); @@ -2344,7 +2348,8 @@ x_143 = lean_ctor_get(x_141, 0); x_144 = lean_ctor_get(x_143, 0); lean_inc(x_144); lean_dec(x_143); -x_145 = lean_environment_main_module(x_144); +x_145 = l_Lean_Environment_mainModule(x_144); +lean_dec(x_144); x_146 = lean_array_size(x_91); x_147 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_146, x_113, x_91); x_148 = l_Lean_mkSepArray(x_147, x_115); @@ -2393,7 +2398,8 @@ lean_dec(x_141); x_163 = lean_ctor_get(x_161, 0); lean_inc(x_163); lean_dec(x_161); -x_164 = lean_environment_main_module(x_163); +x_164 = l_Lean_Environment_mainModule(x_163); +lean_dec(x_163); x_165 = lean_array_size(x_91); x_166 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_165, x_113, x_91); x_167 = l_Lean_mkSepArray(x_166, x_115); @@ -2445,7 +2451,8 @@ lean_dec(x_123); x_183 = lean_ctor_get(x_181, 0); lean_inc(x_183); lean_dec(x_181); -x_184 = lean_environment_main_module(x_183); +x_184 = l_Lean_Environment_mainModule(x_183); +lean_dec(x_183); x_185 = lean_array_size(x_85); x_186 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_185, x_113, x_85); x_187 = l_Lean_mkSepArray(x_186, x_115); @@ -2491,7 +2498,8 @@ if (lean_is_exclusive(x_197)) { x_201 = lean_ctor_get(x_198, 0); lean_inc(x_201); lean_dec(x_198); -x_202 = lean_environment_main_module(x_201); +x_202 = l_Lean_Environment_mainModule(x_201); +lean_dec(x_201); x_203 = lean_array_size(x_91); x_204 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_203, x_113, x_91); x_205 = l_Lean_mkSepArray(x_204, x_115); @@ -2588,7 +2596,8 @@ if (lean_is_exclusive(x_234)) { x_238 = lean_ctor_get(x_235, 0); lean_inc(x_238); lean_dec(x_235); -x_239 = lean_environment_main_module(x_238); +x_239 = l_Lean_Environment_mainModule(x_238); +lean_dec(x_238); x_240 = lean_array_size(x_85); x_241 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_240, x_224, x_85); x_242 = l_Lean_mkSepArray(x_241, x_226); @@ -2634,7 +2643,8 @@ if (lean_is_exclusive(x_252)) { x_256 = lean_ctor_get(x_253, 0); lean_inc(x_256); lean_dec(x_253); -x_257 = lean_environment_main_module(x_256); +x_257 = l_Lean_Environment_mainModule(x_256); +lean_dec(x_256); x_258 = lean_array_size(x_91); x_259 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_258, x_224, x_91); x_260 = l_Lean_mkSepArray(x_259, x_226); @@ -2725,7 +2735,8 @@ lean_dec(x_93); x_281 = lean_ctor_get(x_279, 0); lean_inc(x_281); lean_dec(x_279); -x_282 = lean_environment_main_module(x_281); +x_282 = l_Lean_Environment_mainModule(x_281); +lean_dec(x_281); x_283 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_92); x_284 = l_Lean_addMacroScope(x_282, x_283, x_92); @@ -2809,7 +2820,8 @@ if (lean_is_exclusive(x_308)) { x_312 = lean_ctor_get(x_309, 0); lean_inc(x_312); lean_dec(x_309); -x_313 = lean_environment_main_module(x_312); +x_313 = l_Lean_Environment_mainModule(x_312); +lean_dec(x_312); x_314 = lean_array_size(x_85); x_315 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_314, x_297, x_85); x_316 = l_Lean_mkSepArray(x_315, x_299); @@ -2855,7 +2867,8 @@ if (lean_is_exclusive(x_326)) { x_330 = lean_ctor_get(x_327, 0); lean_inc(x_330); lean_dec(x_327); -x_331 = lean_environment_main_module(x_330); +x_331 = l_Lean_Environment_mainModule(x_330); +lean_dec(x_330); x_332 = lean_array_size(x_91); x_333 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_332, x_297, x_91); x_334 = l_Lean_mkSepArray(x_333, x_299); @@ -2966,7 +2979,8 @@ if (lean_is_exclusive(x_358)) { x_362 = lean_ctor_get(x_359, 0); lean_inc(x_362); lean_dec(x_359); -x_363 = lean_environment_main_module(x_362); +x_363 = l_Lean_Environment_mainModule(x_362); +lean_dec(x_362); x_364 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_357); x_365 = l_Lean_addMacroScope(x_363, x_364, x_357); @@ -3055,7 +3069,8 @@ if (lean_is_exclusive(x_389)) { x_393 = lean_ctor_get(x_390, 0); lean_inc(x_393); lean_dec(x_390); -x_394 = lean_environment_main_module(x_393); +x_394 = l_Lean_Environment_mainModule(x_393); +lean_dec(x_393); x_395 = lean_array_size(x_85); x_396 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_395, x_378, x_85); x_397 = l_Lean_mkSepArray(x_396, x_380); @@ -3101,7 +3116,8 @@ if (lean_is_exclusive(x_407)) { x_411 = lean_ctor_get(x_408, 0); lean_inc(x_411); lean_dec(x_408); -x_412 = lean_environment_main_module(x_411); +x_412 = l_Lean_Environment_mainModule(x_411); +lean_dec(x_411); x_413 = lean_array_size(x_356); x_414 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_413, x_378, x_356); x_415 = l_Lean_mkSepArray(x_414, x_380); @@ -3248,7 +3264,8 @@ if (lean_is_exclusive(x_448)) { x_452 = lean_ctor_get(x_449, 0); lean_inc(x_452); lean_dec(x_449); -x_453 = lean_environment_main_module(x_452); +x_453 = l_Lean_Environment_mainModule(x_452); +lean_dec(x_452); x_454 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_447); x_455 = l_Lean_addMacroScope(x_453, x_454, x_447); @@ -3337,7 +3354,8 @@ if (lean_is_exclusive(x_479)) { x_483 = lean_ctor_get(x_480, 0); lean_inc(x_483); lean_dec(x_480); -x_484 = lean_environment_main_module(x_483); +x_484 = l_Lean_Environment_mainModule(x_483); +lean_dec(x_483); x_485 = lean_array_size(x_440); x_486 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_485, x_468, x_440); x_487 = l_Lean_mkSepArray(x_486, x_470); @@ -3383,7 +3401,8 @@ if (lean_is_exclusive(x_497)) { x_501 = lean_ctor_get(x_498, 0); lean_inc(x_501); lean_dec(x_498); -x_502 = lean_environment_main_module(x_501); +x_502 = l_Lean_Environment_mainModule(x_501); +lean_dec(x_501); x_503 = lean_array_size(x_446); x_504 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_503, x_468, x_446); x_505 = l_Lean_mkSepArray(x_504, x_470); @@ -3567,7 +3586,8 @@ if (lean_is_exclusive(x_547)) { x_551 = lean_ctor_get(x_548, 0); lean_inc(x_551); lean_dec(x_548); -x_552 = lean_environment_main_module(x_551); +x_552 = l_Lean_Environment_mainModule(x_551); +lean_dec(x_551); x_553 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_546); x_554 = l_Lean_addMacroScope(x_552, x_553, x_546); @@ -3656,7 +3676,8 @@ if (lean_is_exclusive(x_578)) { x_582 = lean_ctor_get(x_579, 0); lean_inc(x_582); lean_dec(x_579); -x_583 = lean_environment_main_module(x_582); +x_583 = l_Lean_Environment_mainModule(x_582); +lean_dec(x_582); x_584 = lean_array_size(x_539); x_585 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_584, x_567, x_539); x_586 = l_Lean_mkSepArray(x_585, x_569); @@ -3702,7 +3723,8 @@ if (lean_is_exclusive(x_596)) { x_600 = lean_ctor_get(x_597, 0); lean_inc(x_600); lean_dec(x_597); -x_601 = lean_environment_main_module(x_600); +x_601 = l_Lean_Environment_mainModule(x_600); +lean_dec(x_600); x_602 = lean_array_size(x_545); x_603 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_602, x_567, x_545); x_604 = l_Lean_mkSepArray(x_603, x_569); @@ -3924,7 +3946,8 @@ if (lean_is_exclusive(x_660)) { x_664 = lean_ctor_get(x_661, 0); lean_inc(x_664); lean_dec(x_661); -x_665 = lean_environment_main_module(x_664); +x_665 = l_Lean_Environment_mainModule(x_664); +lean_dec(x_664); x_666 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_659); x_667 = l_Lean_addMacroScope(x_665, x_666, x_659); @@ -4013,7 +4036,8 @@ if (lean_is_exclusive(x_691)) { x_695 = lean_ctor_get(x_692, 0); lean_inc(x_695); lean_dec(x_692); -x_696 = lean_environment_main_module(x_695); +x_696 = l_Lean_Environment_mainModule(x_695); +lean_dec(x_695); x_697 = lean_array_size(x_652); x_698 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_697, x_680, x_652); x_699 = l_Lean_mkSepArray(x_698, x_682); @@ -4059,7 +4083,8 @@ if (lean_is_exclusive(x_709)) { x_713 = lean_ctor_get(x_710, 0); lean_inc(x_713); lean_dec(x_710); -x_714 = lean_environment_main_module(x_713); +x_714 = l_Lean_Environment_mainModule(x_713); +lean_dec(x_713); x_715 = lean_array_size(x_658); x_716 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_715, x_680, x_658); x_717 = l_Lean_mkSepArray(x_716, x_682); @@ -4304,7 +4329,8 @@ if (lean_is_exclusive(x_781)) { x_785 = lean_ctor_get(x_782, 0); lean_inc(x_785); lean_dec(x_782); -x_786 = lean_environment_main_module(x_785); +x_786 = l_Lean_Environment_mainModule(x_785); +lean_dec(x_785); x_787 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_780); x_788 = l_Lean_addMacroScope(x_786, x_787, x_780); @@ -4393,7 +4419,8 @@ if (lean_is_exclusive(x_812)) { x_816 = lean_ctor_get(x_813, 0); lean_inc(x_816); lean_dec(x_813); -x_817 = lean_environment_main_module(x_816); +x_817 = l_Lean_Environment_mainModule(x_816); +lean_dec(x_816); x_818 = lean_array_size(x_773); x_819 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_818, x_801, x_773); x_820 = l_Lean_mkSepArray(x_819, x_803); @@ -4439,7 +4466,8 @@ if (lean_is_exclusive(x_830)) { x_834 = lean_ctor_get(x_831, 0); lean_inc(x_834); lean_dec(x_831); -x_835 = lean_environment_main_module(x_834); +x_835 = l_Lean_Environment_mainModule(x_834); +lean_dec(x_834); x_836 = lean_array_size(x_779); x_837 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_836, x_801, x_779); x_838 = l_Lean_mkSepArray(x_837, x_803); @@ -4699,7 +4727,8 @@ if (lean_is_exclusive(x_905)) { x_909 = lean_ctor_get(x_906, 0); lean_inc(x_909); lean_dec(x_906); -x_910 = lean_environment_main_module(x_909); +x_910 = l_Lean_Environment_mainModule(x_909); +lean_dec(x_909); x_911 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_904); x_912 = l_Lean_addMacroScope(x_910, x_911, x_904); @@ -4788,7 +4817,8 @@ if (lean_is_exclusive(x_936)) { x_940 = lean_ctor_get(x_937, 0); lean_inc(x_940); lean_dec(x_937); -x_941 = lean_environment_main_module(x_940); +x_941 = l_Lean_Environment_mainModule(x_940); +lean_dec(x_940); x_942 = lean_array_size(x_897); x_943 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_942, x_925, x_897); x_944 = l_Lean_mkSepArray(x_943, x_927); @@ -4834,7 +4864,8 @@ if (lean_is_exclusive(x_954)) { x_958 = lean_ctor_get(x_955, 0); lean_inc(x_958); lean_dec(x_955); -x_959 = lean_environment_main_module(x_958); +x_959 = l_Lean_Environment_mainModule(x_958); +lean_dec(x_958); x_960 = lean_array_size(x_903); x_961 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_960, x_925, x_903); x_962 = l_Lean_mkSepArray(x_961, x_927); @@ -5173,7 +5204,8 @@ if (lean_is_exclusive(x_1045)) { x_1049 = lean_ctor_get(x_1046, 0); lean_inc(x_1049); lean_dec(x_1046); -x_1050 = lean_environment_main_module(x_1049); +x_1050 = l_Lean_Environment_mainModule(x_1049); +lean_dec(x_1049); x_1051 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_1044); x_1052 = l_Lean_addMacroScope(x_1050, x_1051, x_1044); @@ -5262,7 +5294,8 @@ if (lean_is_exclusive(x_1076)) { x_1080 = lean_ctor_get(x_1077, 0); lean_inc(x_1080); lean_dec(x_1077); -x_1081 = lean_environment_main_module(x_1080); +x_1081 = l_Lean_Environment_mainModule(x_1080); +lean_dec(x_1080); x_1082 = lean_array_size(x_1037); x_1083 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_1082, x_1065, x_1037); x_1084 = l_Lean_mkSepArray(x_1083, x_1067); @@ -5308,7 +5341,8 @@ if (lean_is_exclusive(x_1094)) { x_1098 = lean_ctor_get(x_1095, 0); lean_inc(x_1098); lean_dec(x_1095); -x_1099 = lean_environment_main_module(x_1098); +x_1099 = l_Lean_Environment_mainModule(x_1098); +lean_dec(x_1098); x_1100 = lean_array_size(x_1043); x_1101 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_1100, x_1065, x_1043); x_1102 = l_Lean_mkSepArray(x_1101, x_1067); @@ -5665,7 +5699,8 @@ if (lean_is_exclusive(x_1188)) { x_1192 = lean_ctor_get(x_1189, 0); lean_inc(x_1192); lean_dec(x_1189); -x_1193 = lean_environment_main_module(x_1192); +x_1193 = l_Lean_Environment_mainModule(x_1192); +lean_dec(x_1192); x_1194 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_1187); x_1195 = l_Lean_addMacroScope(x_1193, x_1194, x_1187); @@ -5754,7 +5789,8 @@ if (lean_is_exclusive(x_1219)) { x_1223 = lean_ctor_get(x_1220, 0); lean_inc(x_1223); lean_dec(x_1220); -x_1224 = lean_environment_main_module(x_1223); +x_1224 = l_Lean_Environment_mainModule(x_1223); +lean_dec(x_1223); x_1225 = lean_array_size(x_1180); x_1226 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_1225, x_1208, x_1180); x_1227 = l_Lean_mkSepArray(x_1226, x_1210); @@ -5800,7 +5836,8 @@ if (lean_is_exclusive(x_1237)) { x_1241 = lean_ctor_get(x_1238, 0); lean_inc(x_1241); lean_dec(x_1238); -x_1242 = lean_environment_main_module(x_1241); +x_1242 = l_Lean_Environment_mainModule(x_1241); +lean_dec(x_1241); x_1243 = lean_array_size(x_1186); x_1244 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_1243, x_1208, x_1186); x_1245 = l_Lean_mkSepArray(x_1244, x_1210); @@ -6181,7 +6218,8 @@ if (lean_is_exclusive(x_1338)) { x_1342 = lean_ctor_get(x_1339, 0); lean_inc(x_1342); lean_dec(x_1339); -x_1343 = lean_environment_main_module(x_1342); +x_1343 = l_Lean_Environment_mainModule(x_1342); +lean_dec(x_1342); x_1344 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__13; lean_inc(x_1337); x_1345 = l_Lean_addMacroScope(x_1343, x_1344, x_1337); @@ -6270,7 +6308,8 @@ if (lean_is_exclusive(x_1369)) { x_1373 = lean_ctor_get(x_1370, 0); lean_inc(x_1373); lean_dec(x_1370); -x_1374 = lean_environment_main_module(x_1373); +x_1374 = l_Lean_Environment_mainModule(x_1373); +lean_dec(x_1373); x_1375 = lean_array_size(x_1330); x_1376 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_1375, x_1358, x_1330); x_1377 = l_Lean_mkSepArray(x_1376, x_1360); @@ -6316,7 +6355,8 @@ if (lean_is_exclusive(x_1387)) { x_1391 = lean_ctor_get(x_1388, 0); lean_inc(x_1391); lean_dec(x_1388); -x_1392 = lean_environment_main_module(x_1391); +x_1392 = l_Lean_Environment_mainModule(x_1391); +lean_dec(x_1391); x_1393 = lean_array_size(x_1336); x_1394 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_1393, x_1358, x_1336); x_1395 = l_Lean_mkSepArray(x_1394, x_1360); @@ -7455,7 +7495,8 @@ x_127 = lean_ctor_get(x_125, 0); x_128 = lean_ctor_get(x_127, 0); lean_inc(x_128); lean_dec(x_127); -x_129 = lean_environment_main_module(x_128); +x_129 = l_Lean_Environment_mainModule(x_128); +lean_dec(x_128); x_130 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13; x_131 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__9; lean_inc(x_123); @@ -7548,7 +7589,8 @@ lean_dec(x_125); x_171 = lean_ctor_get(x_169, 0); lean_inc(x_171); lean_dec(x_169); -x_172 = lean_environment_main_module(x_171); +x_172 = l_Lean_Environment_mainModule(x_171); +lean_dec(x_171); x_173 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13; x_174 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__9; lean_inc(x_123); @@ -7669,7 +7711,8 @@ x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec(x_22); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13; x_26 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__9; lean_inc(x_18); @@ -7772,7 +7815,8 @@ lean_dec(x_20); x_71 = lean_ctor_get(x_69, 0); lean_inc(x_71); lean_dec(x_69); -x_72 = lean_environment_main_module(x_71); +x_72 = l_Lean_Environment_mainModule(x_71); +lean_dec(x_71); x_73 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13; x_74 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__9; lean_inc(x_18); @@ -8471,7 +8515,8 @@ x_24 = lean_ctor_get(x_22, 0); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); lean_dec(x_24); -x_26 = lean_environment_main_module(x_25); +x_26 = l_Lean_Environment_mainModule(x_25); +lean_dec(x_25); x_27 = l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__1; lean_inc(x_20); x_28 = lean_alloc_ctor(2, 2, 0); @@ -8536,7 +8581,8 @@ lean_dec(x_22); x_52 = lean_ctor_get(x_50, 0); lean_inc(x_52); lean_dec(x_50); -x_53 = lean_environment_main_module(x_52); +x_53 = l_Lean_Environment_mainModule(x_52); +lean_dec(x_52); x_54 = l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__1; lean_inc(x_20); x_55 = lean_alloc_ctor(2, 2, 0); @@ -9237,7 +9283,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -9283,7 +9329,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -9336,6 +9382,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_Deriving_Ord_mkOrdInstanceHandler___spec__3(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -9554,6 +9601,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_Deriving_Ord_mkOrdInstanceHandler___spec__3(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/Repr.c b/stage0/stdlib/Lean/Elab/Deriving/Repr.c index 5893b2e2fc46..6db17dc893cc 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Repr.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Repr.c @@ -65,7 +65,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2(l static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__2; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__15; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__1___closed__1; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__19; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___lambda__1___boxed(lean_object**); @@ -245,7 +245,7 @@ static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyFor LEAN_EXPORT uint8_t l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1(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*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__12; static lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__3; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__1; @@ -615,7 +615,8 @@ x_20 = lean_ctor_get(x_18, 0); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); lean_dec(x_20); -x_22 = lean_environment_main_module(x_21); +x_22 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_23 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__8; lean_inc(x_16); x_24 = lean_alloc_ctor(2, 2, 0); @@ -709,7 +710,8 @@ lean_dec(x_18); x_57 = lean_ctor_get(x_55, 0); lean_inc(x_57); lean_dec(x_55); -x_58 = lean_environment_main_module(x_57); +x_58 = l_Lean_Environment_mainModule(x_57); +lean_dec(x_57); x_59 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__8; lean_inc(x_16); x_60 = lean_alloc_ctor(2, 2, 0); @@ -1485,7 +1487,8 @@ x_31 = lean_ctor_get(x_29, 0); x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); lean_dec(x_31); -x_33 = lean_environment_main_module(x_32); +x_33 = l_Lean_Environment_mainModule(x_32); +lean_dec(x_32); x_34 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_27); x_35 = lean_alloc_ctor(2, 2, 0); @@ -1634,7 +1637,8 @@ lean_dec(x_29); x_93 = lean_ctor_get(x_91, 0); lean_inc(x_93); lean_dec(x_91); -x_94 = lean_environment_main_module(x_93); +x_94 = l_Lean_Environment_mainModule(x_93); +lean_dec(x_93); x_95 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_27); x_96 = lean_alloc_ctor(2, 2, 0); @@ -2023,7 +2027,8 @@ x_45 = lean_ctor_get(x_42, 1); x_46 = lean_ctor_get(x_44, 0); lean_inc(x_46); lean_dec(x_44); -x_47 = lean_environment_main_module(x_46); +x_47 = l_Lean_Environment_mainModule(x_46); +lean_dec(x_46); x_48 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_40); lean_ctor_set_tag(x_42, 2); @@ -2179,7 +2184,8 @@ lean_dec(x_42); x_85 = lean_ctor_get(x_83, 0); lean_inc(x_85); lean_dec(x_83); -x_86 = lean_environment_main_module(x_85); +x_86 = l_Lean_Environment_mainModule(x_85); +lean_dec(x_85); x_87 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_40); x_88 = lean_alloc_ctor(2, 2, 0); @@ -2537,7 +2543,8 @@ x_31 = lean_ctor_get(x_29, 0); x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); lean_dec(x_31); -x_33 = lean_environment_main_module(x_32); +x_33 = l_Lean_Environment_mainModule(x_32); +lean_dec(x_32); x_34 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__4; x_35 = l_Lean_addMacroScope(x_33, x_34, x_28); x_36 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__5; @@ -2589,7 +2596,8 @@ lean_dec(x_29); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__4; x_57 = l_Lean_addMacroScope(x_55, x_56, x_28); x_58 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__5; @@ -2819,7 +2827,8 @@ lean_dec(x_19); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = lean_box(0); x_25 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7; x_26 = l_Lean_addMacroScope(x_23, x_25, x_18); @@ -3494,7 +3503,8 @@ x_42 = lean_ctor_get(x_40, 0); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); lean_dec(x_42); -x_44 = lean_environment_main_module(x_43); +x_44 = l_Lean_Environment_mainModule(x_43); +lean_dec(x_43); x_45 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_38); x_46 = lean_alloc_ctor(2, 2, 0); @@ -3567,7 +3577,8 @@ lean_dec(x_40); x_71 = lean_ctor_get(x_69, 0); lean_inc(x_71); lean_dec(x_69); -x_72 = lean_environment_main_module(x_71); +x_72 = l_Lean_Environment_mainModule(x_71); +lean_dec(x_71); x_73 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_38); x_74 = lean_alloc_ctor(2, 2, 0); @@ -3656,7 +3667,8 @@ x_104 = lean_ctor_get(x_102, 0); x_105 = lean_ctor_get(x_104, 0); lean_inc(x_105); lean_dec(x_104); -x_106 = lean_environment_main_module(x_105); +x_106 = l_Lean_Environment_mainModule(x_105); +lean_dec(x_105); x_107 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_100); x_108 = lean_alloc_ctor(2, 2, 0); @@ -3714,7 +3726,8 @@ lean_dec(x_102); x_129 = lean_ctor_get(x_127, 0); lean_inc(x_129); lean_dec(x_127); -x_130 = lean_environment_main_module(x_129); +x_130 = l_Lean_Environment_mainModule(x_129); +lean_dec(x_129); x_131 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_100); x_132 = lean_alloc_ctor(2, 2, 0); @@ -3790,7 +3803,8 @@ x_177 = lean_ctor_get(x_175, 0); x_178 = lean_ctor_get(x_177, 0); lean_inc(x_178); lean_dec(x_177); -x_179 = lean_environment_main_module(x_178); +x_179 = l_Lean_Environment_mainModule(x_178); +lean_dec(x_178); x_180 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_173); x_181 = lean_alloc_ctor(2, 2, 0); @@ -3853,7 +3867,8 @@ lean_dec(x_175); x_205 = lean_ctor_get(x_203, 0); lean_inc(x_205); lean_dec(x_203); -x_206 = lean_environment_main_module(x_205); +x_206 = l_Lean_Environment_mainModule(x_205); +lean_dec(x_205); x_207 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_173); x_208 = lean_alloc_ctor(2, 2, 0); @@ -4158,7 +4173,8 @@ if (lean_is_exclusive(x_254)) { x_258 = lean_ctor_get(x_255, 0); lean_inc(x_258); lean_dec(x_255); -x_259 = lean_environment_main_module(x_258); +x_259 = l_Lean_Environment_mainModule(x_258); +lean_dec(x_258); x_260 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_252); x_261 = lean_alloc_ctor(2, 2, 0); @@ -4257,7 +4273,8 @@ if (lean_is_exclusive(x_289)) { x_293 = lean_ctor_get(x_290, 0); lean_inc(x_293); lean_dec(x_290); -x_294 = lean_environment_main_module(x_293); +x_294 = l_Lean_Environment_mainModule(x_293); +lean_dec(x_293); x_295 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_287); x_296 = lean_alloc_ctor(2, 2, 0); @@ -4343,7 +4360,8 @@ if (lean_is_exclusive(x_339)) { x_343 = lean_ctor_get(x_340, 0); lean_inc(x_343); lean_dec(x_340); -x_344 = lean_environment_main_module(x_343); +x_344 = l_Lean_Environment_mainModule(x_343); +lean_dec(x_343); x_345 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__2___closed__5; lean_inc(x_337); x_346 = lean_alloc_ctor(2, 2, 0); @@ -5426,7 +5444,8 @@ x_38 = lean_ctor_get(x_35, 1); x_39 = lean_ctor_get(x_37, 0); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); +x_40 = l_Lean_Environment_mainModule(x_39); +lean_dec(x_39); x_41 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__6; lean_inc(x_34); x_42 = l_Lean_addMacroScope(x_40, x_41, x_34); @@ -5507,7 +5526,8 @@ x_74 = lean_ctor_get(x_72, 0); x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); lean_dec(x_74); -x_76 = lean_environment_main_module(x_75); +x_76 = l_Lean_Environment_mainModule(x_75); +lean_dec(x_75); x_77 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; lean_inc(x_33); lean_ctor_set_tag(x_54, 2); @@ -5681,7 +5701,8 @@ lean_dec(x_72); x_148 = lean_ctor_get(x_146, 0); lean_inc(x_148); lean_dec(x_146); -x_149 = lean_environment_main_module(x_148); +x_149 = l_Lean_Environment_mainModule(x_148); +lean_dec(x_148); x_150 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; lean_inc(x_33); lean_ctor_set_tag(x_54, 2); @@ -5890,7 +5911,8 @@ if (lean_is_exclusive(x_231)) { x_235 = lean_ctor_get(x_232, 0); lean_inc(x_235); lean_dec(x_232); -x_236 = lean_environment_main_module(x_235); +x_236 = l_Lean_Environment_mainModule(x_235); +lean_dec(x_235); x_237 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; lean_inc(x_33); lean_ctor_set_tag(x_54, 2); @@ -6121,7 +6143,8 @@ if (lean_is_exclusive(x_322)) { x_326 = lean_ctor_get(x_323, 0); lean_inc(x_326); lean_dec(x_323); -x_327 = lean_environment_main_module(x_326); +x_327 = l_Lean_Environment_mainModule(x_326); +lean_dec(x_326); x_328 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; lean_inc(x_33); x_329 = lean_alloc_ctor(2, 2, 0); @@ -6331,7 +6354,8 @@ lean_dec(x_35); x_405 = lean_ctor_get(x_403, 0); lean_inc(x_405); lean_dec(x_403); -x_406 = lean_environment_main_module(x_405); +x_406 = l_Lean_Environment_mainModule(x_405); +lean_dec(x_405); x_407 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__6; lean_inc(x_34); x_408 = l_Lean_addMacroScope(x_406, x_407, x_34); @@ -6434,7 +6458,8 @@ if (lean_is_exclusive(x_439)) { x_443 = lean_ctor_get(x_440, 0); lean_inc(x_443); lean_dec(x_440); -x_444 = lean_environment_main_module(x_443); +x_444 = l_Lean_Environment_mainModule(x_443); +lean_dec(x_443); x_445 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; lean_inc(x_33); if (lean_is_scalar(x_425)) { @@ -6680,7 +6705,8 @@ if (lean_is_exclusive(x_532)) { x_536 = lean_ctor_get(x_533, 0); lean_inc(x_536); lean_dec(x_533); -x_537 = lean_environment_main_module(x_536); +x_537 = l_Lean_Environment_mainModule(x_536); +lean_dec(x_536); x_538 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__6; lean_inc(x_531); x_539 = l_Lean_addMacroScope(x_537, x_538, x_531); @@ -6787,7 +6813,8 @@ if (lean_is_exclusive(x_570)) { x_574 = lean_ctor_get(x_571, 0); lean_inc(x_574); lean_dec(x_571); -x_575 = lean_environment_main_module(x_574); +x_575 = l_Lean_Environment_mainModule(x_574); +lean_dec(x_574); x_576 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; lean_inc(x_530); if (lean_is_scalar(x_556)) { @@ -8053,7 +8080,8 @@ x_21 = lean_ctor_get(x_19, 0); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec(x_21); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__10; x_25 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__23; lean_inc(x_17); @@ -8146,7 +8174,8 @@ lean_dec(x_19); x_65 = lean_ctor_get(x_63, 0); lean_inc(x_65); lean_dec(x_63); -x_66 = lean_environment_main_module(x_65); +x_66 = l_Lean_Environment_mainModule(x_65); +lean_dec(x_65); x_67 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__10; x_68 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__23; lean_inc(x_17); @@ -8252,7 +8281,8 @@ x_114 = lean_ctor_get(x_112, 0); x_115 = lean_ctor_get(x_114, 0); lean_inc(x_115); lean_dec(x_114); -x_116 = lean_environment_main_module(x_115); +x_116 = l_Lean_Environment_mainModule(x_115); +lean_dec(x_115); x_117 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__10; x_118 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__23; lean_inc(x_110); @@ -8355,7 +8385,8 @@ lean_dec(x_112); x_163 = lean_ctor_get(x_161, 0); lean_inc(x_163); lean_dec(x_161); -x_164 = lean_environment_main_module(x_163); +x_164 = l_Lean_Environment_mainModule(x_163); +lean_dec(x_163); x_165 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__10; x_166 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__23; lean_inc(x_110); @@ -9557,7 +9588,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -9603,7 +9634,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -9656,6 +9687,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__3(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -9874,6 +9906,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__3(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/SizeOf.c b/stage0/stdlib/Lean/Elab/Deriving/SizeOf.c index f46b42bbdd08..e6f203a6f195 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/SizeOf.c +++ b/stage0/stdlib/Lean/Elab/Deriving/SizeOf.c @@ -17,7 +17,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_Si lean_object* l_Lean_Meta_mkSizeOfInstances(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_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_SizeOf_mkSizeOfHandler___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_SizeOf_initFn____x40_Lean_Elab_Deriving_SizeOf___hyg_125____closed__1; LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_SizeOf_mkSizeOfHandler___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_SizeOf_mkSizeOfHandler(lean_object*, lean_object*, lean_object*, lean_object*); @@ -122,7 +122,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -168,7 +168,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -221,6 +221,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_Deriving_SizeOf_mkSizeOfHandler___spec__2(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -436,6 +437,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_Deriving_SizeOf_mkSizeOfHandler___spec__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c b/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c index 6338a0682504..d3e0717f3f37 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c +++ b/stage0/stdlib/Lean/Elab/Deriving/ToExpr.c @@ -84,7 +84,7 @@ static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___clos static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__7; static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__20; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkInstanceCmds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkInstanceCmds___spec__3___closed__10; static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__19; @@ -294,7 +294,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_ToExpr_mkToExprBody(lean_object*, lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__32; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__4; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkLocalInstanceLetDecls___spec__1___closed__13; static lean_object* l_Lean_Elab_Deriving_ToExpr_mkAuxFunction___lambda__1___closed__4; @@ -597,7 +597,8 @@ lean_dec(x_16); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_main_module(x_19); +x_20 = l_Lean_Environment_mainModule(x_19); +lean_dec(x_19); x_21 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__9; x_22 = l_Lean_addMacroScope(x_20, x_21, x_15); x_23 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__7; @@ -1259,7 +1260,8 @@ x_23 = lean_ctor_get(x_20, 1); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_19); x_28 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2; @@ -1305,7 +1307,8 @@ lean_dec(x_20); x_45 = lean_ctor_get(x_43, 0); lean_inc(x_45); lean_dec(x_43); -x_46 = lean_environment_main_module(x_45); +x_46 = l_Lean_Environment_mainModule(x_45); +lean_dec(x_45); x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__4; x_48 = l_Lean_addMacroScope(x_46, x_47, x_19); x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__1___closed__2; @@ -1575,7 +1578,8 @@ lean_dec(x_50); x_53 = lean_ctor_get(x_51, 0); lean_inc(x_53); lean_dec(x_51); -x_54 = lean_environment_main_module(x_53); +x_54 = l_Lean_Environment_mainModule(x_53); +lean_dec(x_53); x_55 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; x_56 = l_Lean_addMacroScope(x_54, x_55, x_49); x_57 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; @@ -1621,7 +1625,8 @@ lean_dec(x_71); x_74 = lean_ctor_get(x_72, 0); lean_inc(x_74); lean_dec(x_72); -x_75 = lean_environment_main_module(x_74); +x_75 = l_Lean_Environment_mainModule(x_74); +lean_dec(x_74); x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; x_77 = l_Lean_addMacroScope(x_75, x_76, x_70); x_78 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; @@ -1724,7 +1729,8 @@ lean_dec(x_104); x_107 = lean_ctor_get(x_105, 0); lean_inc(x_107); lean_dec(x_105); -x_108 = lean_environment_main_module(x_107); +x_108 = l_Lean_Environment_mainModule(x_107); +lean_dec(x_107); x_109 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; x_110 = l_Lean_addMacroScope(x_108, x_109, x_103); x_111 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; @@ -1771,7 +1777,8 @@ lean_dec(x_125); x_128 = lean_ctor_get(x_126, 0); lean_inc(x_128); lean_dec(x_126); -x_129 = lean_environment_main_module(x_128); +x_129 = l_Lean_Environment_mainModule(x_128); +lean_dec(x_128); x_130 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; x_131 = l_Lean_addMacroScope(x_129, x_130, x_124); x_132 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; @@ -1922,7 +1929,8 @@ lean_dec(x_167); x_170 = lean_ctor_get(x_168, 0); lean_inc(x_170); lean_dec(x_168); -x_171 = lean_environment_main_module(x_170); +x_171 = l_Lean_Environment_mainModule(x_170); +lean_dec(x_170); x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; x_173 = l_Lean_addMacroScope(x_171, x_172, x_166); x_174 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; @@ -1970,7 +1978,8 @@ lean_dec(x_189); x_192 = lean_ctor_get(x_190, 0); lean_inc(x_192); lean_dec(x_190); -x_193 = lean_environment_main_module(x_192); +x_193 = l_Lean_Environment_mainModule(x_192); +lean_dec(x_192); x_194 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; x_195 = l_Lean_addMacroScope(x_193, x_194, x_188); x_196 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; @@ -2298,7 +2307,8 @@ x_35 = lean_ctor_get(x_32, 1); x_36 = lean_ctor_get(x_34, 0); lean_inc(x_36); lean_dec(x_34); -x_37 = lean_environment_main_module(x_36); +x_37 = l_Lean_Environment_mainModule(x_36); +lean_dec(x_36); x_38 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; x_39 = l_Lean_addMacroScope(x_37, x_38, x_31); x_40 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; @@ -2399,7 +2409,8 @@ lean_dec(x_32); x_77 = lean_ctor_get(x_75, 0); lean_inc(x_77); lean_dec(x_75); -x_78 = lean_environment_main_module(x_77); +x_78 = l_Lean_Environment_mainModule(x_77); +lean_dec(x_77); x_79 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; x_80 = l_Lean_addMacroScope(x_78, x_79, x_31); x_81 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; @@ -2519,7 +2530,8 @@ if (lean_is_exclusive(x_122)) { x_126 = lean_ctor_get(x_123, 0); lean_inc(x_126); lean_dec(x_123); -x_127 = lean_environment_main_module(x_126); +x_127 = l_Lean_Environment_mainModule(x_126); +lean_dec(x_126); x_128 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; x_129 = l_Lean_addMacroScope(x_127, x_128, x_121); x_130 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; @@ -2810,7 +2822,8 @@ x_23 = lean_ctor_get(x_20, 1); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; lean_inc(x_18); lean_ctor_set_tag(x_20, 2); @@ -2847,7 +2860,8 @@ lean_dec(x_20); x_40 = lean_ctor_get(x_38, 0); lean_inc(x_40); lean_dec(x_38); -x_41 = lean_environment_main_module(x_40); +x_41 = l_Lean_Environment_mainModule(x_40); +lean_dec(x_40); x_42 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__18; lean_inc(x_18); x_43 = lean_alloc_ctor(2, 2, 0); @@ -3346,7 +3360,8 @@ x_64 = lean_ctor_get(x_61, 1); x_65 = lean_ctor_get(x_63, 0); lean_inc(x_65); lean_dec(x_63); -x_66 = lean_environment_main_module(x_65); +x_66 = l_Lean_Environment_mainModule(x_65); +lean_dec(x_65); x_67 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; x_68 = l_Lean_addMacroScope(x_66, x_67, x_60); x_69 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; @@ -3382,7 +3397,8 @@ lean_dec(x_61); x_80 = lean_ctor_get(x_78, 0); lean_inc(x_80); lean_dec(x_78); -x_81 = lean_environment_main_module(x_80); +x_81 = l_Lean_Environment_mainModule(x_80); +lean_dec(x_80); x_82 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; x_83 = l_Lean_addMacroScope(x_81, x_82, x_60); x_84 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; @@ -3427,7 +3443,8 @@ x_99 = lean_ctor_get(x_96, 1); x_100 = lean_ctor_get(x_98, 0); lean_inc(x_100); lean_dec(x_98); -x_101 = lean_environment_main_module(x_100); +x_101 = l_Lean_Environment_mainModule(x_100); +lean_dec(x_100); x_102 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; x_103 = l_Lean_addMacroScope(x_101, x_102, x_95); x_104 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; @@ -3463,7 +3480,8 @@ lean_dec(x_96); x_115 = lean_ctor_get(x_113, 0); lean_inc(x_115); lean_dec(x_113); -x_116 = lean_environment_main_module(x_115); +x_116 = l_Lean_Environment_mainModule(x_115); +lean_dec(x_115); x_117 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; x_118 = l_Lean_addMacroScope(x_116, x_117, x_95); x_119 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; @@ -3718,7 +3736,8 @@ if (lean_is_exclusive(x_188)) { x_192 = lean_ctor_get(x_189, 0); lean_inc(x_192); lean_dec(x_189); -x_193 = lean_environment_main_module(x_192); +x_193 = l_Lean_Environment_mainModule(x_192); +lean_dec(x_192); x_194 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; x_195 = l_Lean_addMacroScope(x_193, x_194, x_187); x_196 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; @@ -3772,7 +3791,8 @@ if (lean_is_exclusive(x_208)) { x_212 = lean_ctor_get(x_209, 0); lean_inc(x_212); lean_dec(x_209); -x_213 = lean_environment_main_module(x_212); +x_213 = l_Lean_Environment_mainModule(x_212); +lean_dec(x_212); x_214 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; x_215 = l_Lean_addMacroScope(x_213, x_214, x_207); x_216 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; @@ -4074,7 +4094,8 @@ x_55 = lean_ctor_get(x_52, 1); x_56 = lean_ctor_get(x_54, 0); lean_inc(x_56); lean_dec(x_54); -x_57 = lean_environment_main_module(x_56); +x_57 = l_Lean_Environment_mainModule(x_56); +lean_dec(x_56); x_58 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; x_59 = l_Lean_addMacroScope(x_57, x_58, x_51); x_60 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; @@ -4110,7 +4131,8 @@ lean_dec(x_52); x_71 = lean_ctor_get(x_69, 0); lean_inc(x_71); lean_dec(x_69); -x_72 = lean_environment_main_module(x_71); +x_72 = l_Lean_Environment_mainModule(x_71); +lean_dec(x_71); x_73 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__3; x_74 = l_Lean_addMacroScope(x_72, x_73, x_51); x_75 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__2; @@ -4160,7 +4182,8 @@ x_93 = lean_ctor_get(x_90, 1); x_94 = lean_ctor_get(x_92, 0); lean_inc(x_94); lean_dec(x_92); -x_95 = lean_environment_main_module(x_94); +x_95 = l_Lean_Environment_mainModule(x_94); +lean_dec(x_94); x_96 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; x_97 = l_Lean_addMacroScope(x_95, x_96, x_89); x_98 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; @@ -4196,7 +4219,8 @@ lean_dec(x_90); x_109 = lean_ctor_get(x_107, 0); lean_inc(x_109); lean_dec(x_107); -x_110 = lean_environment_main_module(x_109); +x_110 = l_Lean_Environment_mainModule(x_109); +lean_dec(x_109); x_111 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__10; x_112 = l_Lean_addMacroScope(x_110, x_111, x_89); x_113 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___spec__2___closed__9; @@ -4548,7 +4572,8 @@ x_121 = lean_ctor_get(x_118, 1); x_122 = lean_ctor_get(x_120, 0); lean_inc(x_122); lean_dec(x_120); -x_123 = lean_environment_main_module(x_122); +x_123 = l_Lean_Environment_mainModule(x_122); +lean_dec(x_122); x_124 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; x_125 = l_Lean_addMacroScope(x_123, x_124, x_117); x_126 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; @@ -4641,7 +4666,8 @@ lean_dec(x_118); x_157 = lean_ctor_get(x_155, 0); lean_inc(x_157); lean_dec(x_155); -x_158 = lean_environment_main_module(x_157); +x_158 = l_Lean_Environment_mainModule(x_157); +lean_dec(x_157); x_159 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; x_160 = l_Lean_addMacroScope(x_158, x_159, x_117); x_161 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; @@ -4929,7 +4955,8 @@ if (lean_is_exclusive(x_229)) { x_233 = lean_ctor_get(x_230, 0); lean_inc(x_233); lean_dec(x_230); -x_234 = lean_environment_main_module(x_233); +x_234 = l_Lean_Environment_mainModule(x_233); +lean_dec(x_233); x_235 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; x_236 = l_Lean_addMacroScope(x_234, x_235, x_228); x_237 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; @@ -5228,7 +5255,8 @@ if (lean_is_exclusive(x_325)) { x_329 = lean_ctor_get(x_326, 0); lean_inc(x_329); lean_dec(x_326); -x_330 = lean_environment_main_module(x_329); +x_330 = l_Lean_Environment_mainModule(x_329); +lean_dec(x_329); x_331 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; x_332 = l_Lean_addMacroScope(x_330, x_331, x_324); x_333 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; @@ -5599,7 +5627,8 @@ if (lean_is_exclusive(x_434)) { x_438 = lean_ctor_get(x_435, 0); lean_inc(x_438); lean_dec(x_435); -x_439 = lean_environment_main_module(x_438); +x_439 = l_Lean_Environment_mainModule(x_438); +lean_dec(x_438); x_440 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__5; x_441 = l_Lean_addMacroScope(x_439, x_440, x_433); x_442 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__3; @@ -7217,7 +7246,8 @@ x_80 = lean_ctor_get(x_77, 1); x_81 = lean_ctor_get(x_79, 0); lean_inc(x_81); lean_dec(x_79); -x_82 = lean_environment_main_module(x_81); +x_82 = l_Lean_Environment_mainModule(x_81); +lean_dec(x_81); x_83 = lean_mk_syntax_ident(x_68); x_84 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; x_85 = l_Array_append___rarg(x_84, x_52); @@ -7373,7 +7403,8 @@ lean_dec(x_77); x_146 = lean_ctor_get(x_144, 0); lean_inc(x_146); lean_dec(x_144); -x_147 = lean_environment_main_module(x_146); +x_147 = l_Lean_Environment_mainModule(x_146); +lean_dec(x_146); x_148 = lean_mk_syntax_ident(x_68); x_149 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; x_150 = l_Array_append___rarg(x_149, x_52); @@ -7606,7 +7637,8 @@ if (lean_is_exclusive(x_223)) { x_227 = lean_ctor_get(x_224, 0); lean_inc(x_227); lean_dec(x_224); -x_228 = lean_environment_main_module(x_227); +x_228 = l_Lean_Environment_mainModule(x_227); +lean_dec(x_227); x_229 = lean_mk_syntax_ident(x_214); x_230 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; x_231 = l_Array_append___rarg(x_230, x_52); @@ -7859,7 +7891,8 @@ if (lean_is_exclusive(x_310)) { x_314 = lean_ctor_get(x_311, 0); lean_inc(x_314); lean_dec(x_311); -x_315 = lean_environment_main_module(x_314); +x_315 = l_Lean_Environment_mainModule(x_314); +lean_dec(x_314); x_316 = lean_mk_syntax_ident(x_300); x_317 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; x_318 = l_Array_append___rarg(x_317, x_52); @@ -8142,7 +8175,8 @@ if (lean_is_exclusive(x_409)) { x_413 = lean_ctor_get(x_410, 0); lean_inc(x_413); lean_dec(x_410); -x_414 = lean_environment_main_module(x_413); +x_414 = l_Lean_Environment_mainModule(x_413); +lean_dec(x_413); x_415 = lean_mk_syntax_ident(x_399); x_416 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; x_417 = l_Array_append___rarg(x_416, x_384); @@ -8506,7 +8540,8 @@ if (lean_is_exclusive(x_526)) { x_530 = lean_ctor_get(x_527, 0); lean_inc(x_530); lean_dec(x_527); -x_531 = lean_environment_main_module(x_530); +x_531 = l_Lean_Environment_mainModule(x_530); +lean_dec(x_530); x_532 = lean_mk_syntax_ident(x_516); x_533 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; x_534 = l_Array_append___rarg(x_533, x_500); @@ -8923,7 +8958,8 @@ if (lean_is_exclusive(x_652)) { x_656 = lean_ctor_get(x_653, 0); lean_inc(x_656); lean_dec(x_653); -x_657 = lean_environment_main_module(x_656); +x_657 = l_Lean_Environment_mainModule(x_656); +lean_dec(x_656); x_658 = lean_mk_syntax_ident(x_642); x_659 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; x_660 = l_Array_append___rarg(x_659, x_626); @@ -9493,7 +9529,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; lean_inc(x_24); lean_ctor_set_tag(x_26, 2); @@ -9571,7 +9608,8 @@ lean_dec(x_26); x_63 = lean_ctor_get(x_61, 0); lean_inc(x_63); lean_dec(x_61); -x_64 = lean_environment_main_module(x_63); +x_64 = l_Lean_Environment_mainModule(x_63); +lean_dec(x_63); x_65 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; lean_inc(x_24); x_66 = lean_alloc_ctor(2, 2, 0); @@ -9671,7 +9709,8 @@ if (lean_is_exclusive(x_102)) { x_106 = lean_ctor_get(x_103, 0); lean_inc(x_106); lean_dec(x_103); -x_107 = lean_environment_main_module(x_106); +x_107 = l_Lean_Environment_mainModule(x_106); +lean_dec(x_106); x_108 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_ToExpr_mkAuxFunction___spec__1___closed__5; lean_inc(x_100); if (lean_is_scalar(x_105)) { @@ -10539,7 +10578,8 @@ x_39 = lean_ctor_get(x_37, 0); x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); lean_dec(x_39); -x_41 = lean_environment_main_module(x_40); +x_41 = l_Lean_Environment_mainModule(x_40); +lean_dec(x_40); x_42 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_43 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_35); @@ -10633,7 +10673,8 @@ lean_dec(x_37); x_83 = lean_ctor_get(x_81, 0); lean_inc(x_83); lean_dec(x_81); -x_84 = lean_environment_main_module(x_83); +x_84 = l_Lean_Environment_mainModule(x_83); +lean_dec(x_83); x_85 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_86 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_35); @@ -10739,7 +10780,8 @@ x_131 = lean_ctor_get(x_129, 0); x_132 = lean_ctor_get(x_131, 0); lean_inc(x_132); lean_dec(x_131); -x_133 = lean_environment_main_module(x_132); +x_133 = l_Lean_Environment_mainModule(x_132); +lean_dec(x_132); x_134 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_135 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_127); @@ -10843,7 +10885,8 @@ lean_dec(x_129); x_180 = lean_ctor_get(x_178, 0); lean_inc(x_180); lean_dec(x_178); -x_181 = lean_environment_main_module(x_180); +x_181 = l_Lean_Environment_mainModule(x_180); +lean_dec(x_180); x_182 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_183 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_127); @@ -11617,7 +11660,8 @@ x_30 = lean_ctor_get(x_27, 1); x_31 = lean_ctor_get(x_29, 0); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_environment_main_module(x_31); +x_32 = l_Lean_Environment_mainModule(x_31); +lean_dec(x_31); x_33 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; lean_inc(x_25); lean_ctor_set_tag(x_27, 2); @@ -11687,7 +11731,8 @@ lean_dec(x_27); x_61 = lean_ctor_get(x_59, 0); lean_inc(x_61); lean_dec(x_59); -x_62 = lean_environment_main_module(x_61); +x_62 = l_Lean_Environment_mainModule(x_61); +lean_dec(x_61); x_63 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; lean_inc(x_25); x_64 = lean_alloc_ctor(2, 2, 0); @@ -11779,7 +11824,8 @@ if (lean_is_exclusive(x_97)) { x_101 = lean_ctor_get(x_98, 0); lean_inc(x_101); lean_dec(x_98); -x_102 = lean_environment_main_module(x_101); +x_102 = l_Lean_Environment_mainModule(x_101); +lean_dec(x_101); x_103 = l_Lean_Elab_Deriving_ToExpr_mkToTypeExpr___lambda__1___closed__14; lean_inc(x_95); if (lean_is_scalar(x_100)) { @@ -12196,7 +12242,8 @@ x_88 = lean_ctor_get(x_85, 1); x_89 = lean_ctor_get(x_87, 0); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_89); +x_90 = l_Lean_Environment_mainModule(x_89); +lean_dec(x_89); x_91 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_92 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_83); @@ -12356,7 +12403,8 @@ lean_dec(x_85); x_157 = lean_ctor_get(x_155, 0); lean_inc(x_157); lean_dec(x_155); -x_158 = lean_environment_main_module(x_157); +x_158 = l_Lean_Environment_mainModule(x_157); +lean_dec(x_157); x_159 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_160 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_83); @@ -12642,7 +12690,8 @@ if (lean_is_exclusive(x_244)) { x_248 = lean_ctor_get(x_245, 0); lean_inc(x_248); lean_dec(x_245); -x_249 = lean_environment_main_module(x_248); +x_249 = l_Lean_Environment_mainModule(x_248); +lean_dec(x_248); x_250 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_251 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_242); @@ -12952,7 +13001,8 @@ if (lean_is_exclusive(x_341)) { x_345 = lean_ctor_get(x_342, 0); lean_inc(x_345); lean_dec(x_342); -x_346 = lean_environment_main_module(x_345); +x_346 = l_Lean_Environment_mainModule(x_345); +lean_dec(x_345); x_347 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_348 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_339); @@ -13281,7 +13331,8 @@ if (lean_is_exclusive(x_443)) { x_447 = lean_ctor_get(x_444, 0); lean_inc(x_447); lean_dec(x_444); -x_448 = lean_environment_main_module(x_447); +x_448 = l_Lean_Environment_mainModule(x_447); +lean_dec(x_447); x_449 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_450 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_441); @@ -13778,7 +13829,8 @@ if (lean_is_exclusive(x_580)) { x_584 = lean_ctor_get(x_581, 0); lean_inc(x_584); lean_dec(x_581); -x_585 = lean_environment_main_module(x_584); +x_585 = l_Lean_Environment_mainModule(x_584); +lean_dec(x_584); x_586 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_587 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_578); @@ -14333,7 +14385,8 @@ if (lean_is_exclusive(x_727)) { x_731 = lean_ctor_get(x_728, 0); lean_inc(x_731); lean_dec(x_728); -x_732 = lean_environment_main_module(x_731); +x_732 = l_Lean_Environment_mainModule(x_731); +lean_dec(x_731); x_733 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_ToExpr_mkAppNTerm___spec__1___closed__17; x_734 = l_Lean_Elab_Deriving_ToExpr_updateIndType___closed__9; lean_inc(x_725); @@ -15218,7 +15271,7 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; lean_object* x_11; @@ -15264,7 +15317,7 @@ lean_dec(x_5); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_find(x_19, x_1); +x_20 = l_Lean_Environment_find_x3f(x_19, x_1); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; lean_object* x_23; @@ -15317,6 +15370,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_array_uget(x_1, x_2); x_9 = l_Lean_isInductive___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__1(x_8, x_4, x_5, x_6); +lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_unbox(x_10); @@ -15644,6 +15698,7 @@ lean_object* x_5; x_5 = l_Lean_isInductive___at_Lean_Elab_Deriving_ToExpr_mkToExprInstanceHandler___spec__1(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/TypeName.c b/stage0/stdlib/Lean/Elab/Deriving/TypeName.c index 2c9f03eaf4f1..c478a486db1d 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/TypeName.c +++ b/stage0/stdlib/Lean/Elab/Deriving/TypeName.c @@ -34,7 +34,7 @@ lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_ob static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__3___lambda__1___closed__44; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__3___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__3___lambda__1___closed__50; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__3___lambda__1___closed__3; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__3___lambda__1___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Deriving_TypeName_0__Lean_Elab_deriveTypeNameInstance___spec__3___lambda__1___closed__46; @@ -290,8 +290,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -332,8 +331,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Elab/Do.c b/stage0/stdlib/Lean/Elab/Do.c index 1c385d6a933f..9a6618a331b2 100644 --- a/stage0/stdlib/Lean/Elab/Do.c +++ b/stage0/stdlib/Lean/Elab/Do.c @@ -1089,7 +1089,7 @@ static lean_object* l_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___closed__1 static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__17; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___closed__12; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_eraseVars___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkTuple___spec__1___closed__3; @@ -7510,7 +7510,8 @@ lean_dec(x_16); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_main_module(x_19); +x_20 = l_Lean_Environment_mainModule(x_19); +lean_dec(x_19); x_21 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__2___closed__3; x_22 = l_Lean_addMacroScope(x_20, x_21, x_15); x_23 = lean_box(0); @@ -8758,7 +8759,8 @@ lean_dec(x_26); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_29); +x_30 = l_Lean_Environment_mainModule(x_29); +lean_dec(x_29); x_31 = l_Lean_Elab_Term_Do_mkSimpleJmp___closed__5; x_32 = l_Lean_addMacroScope(x_30, x_31, x_25); x_33 = lean_box(0); @@ -9319,7 +9321,8 @@ lean_dec(x_26); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_13); +x_30 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_25); lean_ctor_set(x_31, 1, x_30); @@ -9461,7 +9464,8 @@ lean_dec(x_20); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__2___closed__3; x_26 = l_Lean_addMacroScope(x_24, x_25, x_17); x_27 = lean_box(0); @@ -9718,7 +9722,8 @@ lean_dec(x_15); x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); -x_19 = lean_environment_main_module(x_18); +x_19 = l_Lean_Environment_mainModule(x_18); +lean_dec(x_18); x_20 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__2___closed__3; lean_inc(x_14); x_21 = l_Lean_addMacroScope(x_19, x_20, x_14); @@ -9739,7 +9744,8 @@ lean_dec(x_25); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__9; lean_inc(x_13); x_31 = lean_alloc_ctor(2, 2, 0); @@ -13984,7 +13990,8 @@ lean_dec(x_25); x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_environment_main_module(x_12); +x_29 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_30 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); @@ -14270,7 +14277,8 @@ lean_dec(x_23); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -x_27 = lean_environment_main_module(x_26); +x_27 = l_Lean_Environment_mainModule(x_26); +lean_dec(x_26); x_28 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__2___closed__3; x_29 = l_Lean_addMacroScope(x_27, x_28, x_22); x_30 = lean_box(0); @@ -30484,7 +30492,8 @@ lean_dec(x_26); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_13); +x_30 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_25); lean_ctor_set(x_31, 1, x_30); @@ -33272,7 +33281,8 @@ lean_dec(x_5); x_8 = lean_ctor_get(x_6, 0); lean_inc(x_8); lean_dec(x_6); -x_9 = lean_environment_main_module(x_8); +x_9 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_10 = lean_ctor_get(x_2, 10); lean_inc(x_10); lean_dec(x_2); @@ -34365,7 +34375,8 @@ lean_dec(x_26); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_13); +x_30 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_25); lean_ctor_set(x_31, 1, x_30); @@ -34647,7 +34658,8 @@ lean_dec(x_26); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_13); +x_30 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_25); lean_ctor_set(x_31, 1, x_30); @@ -36462,7 +36474,8 @@ lean_dec(x_40); x_43 = lean_ctor_get(x_41, 0); lean_inc(x_43); lean_dec(x_41); -x_44 = lean_environment_main_module(x_43); +x_44 = l_Lean_Environment_mainModule(x_43); +lean_dec(x_43); x_45 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__2___closed__9; lean_inc(x_39); x_46 = l_Lean_addMacroScope(x_44, x_45, x_39); @@ -36483,7 +36496,8 @@ lean_dec(x_50); x_53 = lean_ctor_get(x_51, 0); lean_inc(x_53); lean_dec(x_51); -x_54 = lean_environment_main_module(x_53); +x_54 = l_Lean_Environment_mainModule(x_53); +lean_dec(x_53); x_55 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_38); x_56 = lean_alloc_ctor(2, 2, 0); @@ -36976,7 +36990,8 @@ lean_dec(x_26); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_13); +x_30 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_25); lean_ctor_set(x_31, 1, x_30); @@ -37311,7 +37326,8 @@ lean_dec(x_43); x_46 = lean_ctor_get(x_44, 0); lean_inc(x_46); lean_dec(x_44); -x_47 = lean_environment_main_module(x_46); +x_47 = l_Lean_Environment_mainModule(x_46); +lean_dec(x_46); x_48 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10___closed__3; x_49 = l_Lean_addMacroScope(x_47, x_48, x_42); x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10___closed__2; @@ -37401,7 +37417,8 @@ lean_dec(x_85); x_88 = lean_ctor_get(x_86, 0); lean_inc(x_88); lean_dec(x_86); -x_89 = lean_environment_main_module(x_88); +x_89 = l_Lean_Environment_mainModule(x_88); +lean_dec(x_88); x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10___closed__10; x_91 = l_Lean_addMacroScope(x_89, x_90, x_84); x_92 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10___closed__7; @@ -37690,7 +37707,8 @@ lean_dec(x_27); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__2___closed__3; x_33 = l_Lean_addMacroScope(x_31, x_32, x_26); x_34 = l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__2___closed__2; @@ -38907,7 +38925,8 @@ lean_dec(x_26); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_13); +x_30 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_25); lean_ctor_set(x_31, 1, x_30); @@ -39594,7 +39613,8 @@ lean_dec(x_20); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_18); x_26 = lean_alloc_ctor(2, 2, 0); @@ -39695,7 +39715,8 @@ lean_dec(x_68); x_71 = lean_ctor_get(x_69, 0); lean_inc(x_71); lean_dec(x_69); -x_72 = lean_environment_main_module(x_71); +x_72 = l_Lean_Environment_mainModule(x_71); +lean_dec(x_71); x_73 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_66); x_74 = lean_alloc_ctor(2, 2, 0); @@ -39878,7 +39899,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_19); x_27 = lean_alloc_ctor(2, 2, 0); @@ -40490,7 +40512,8 @@ lean_dec(x_19); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___spec__5___closed__5; lean_inc(x_18); x_25 = lean_alloc_ctor(2, 2, 0); @@ -40536,7 +40559,8 @@ lean_dec(x_43); x_46 = lean_ctor_get(x_44, 0); lean_inc(x_46); lean_dec(x_44); -x_47 = lean_environment_main_module(x_46); +x_47 = l_Lean_Environment_mainModule(x_46); +lean_dec(x_46); x_48 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_42); x_49 = lean_alloc_ctor(2, 2, 0); @@ -41367,7 +41391,8 @@ lean_dec(x_72); x_75 = lean_ctor_get(x_73, 0); lean_inc(x_75); lean_dec(x_73); -x_76 = lean_environment_main_module(x_75); +x_76 = l_Lean_Environment_mainModule(x_75); +lean_dec(x_75); x_77 = l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__5; lean_inc(x_28); x_78 = l_Lean_addMacroScope(x_76, x_77, x_28); @@ -41404,7 +41429,8 @@ lean_dec(x_86); x_89 = lean_ctor_get(x_87, 0); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_89); +x_90 = l_Lean_Environment_mainModule(x_89); +lean_dec(x_89); x_91 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__3; lean_inc(x_71); x_92 = lean_alloc_ctor(2, 2, 0); @@ -41474,7 +41500,8 @@ lean_dec(x_116); x_119 = lean_ctor_get(x_117, 0); lean_inc(x_119); lean_dec(x_117); -x_120 = lean_environment_main_module(x_119); +x_120 = l_Lean_Environment_mainModule(x_119); +lean_dec(x_119); x_121 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__6; lean_inc(x_71); x_122 = lean_alloc_ctor(2, 2, 0); @@ -41569,7 +41596,8 @@ lean_dec(x_148); x_151 = lean_ctor_get(x_149, 0); lean_inc(x_151); lean_dec(x_149); -x_152 = lean_environment_main_module(x_151); +x_152 = l_Lean_Environment_mainModule(x_151); +lean_dec(x_151); x_153 = l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__5; lean_inc(x_28); x_154 = l_Lean_addMacroScope(x_152, x_153, x_28); @@ -41604,7 +41632,8 @@ lean_dec(x_162); x_165 = lean_ctor_get(x_163, 0); lean_inc(x_165); lean_dec(x_163); -x_166 = lean_environment_main_module(x_165); +x_166 = l_Lean_Environment_mainModule(x_165); +lean_dec(x_165); x_167 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__8; lean_inc(x_28); x_168 = l_Lean_addMacroScope(x_166, x_167, x_28); @@ -41636,7 +41665,8 @@ lean_dec(x_177); x_180 = lean_ctor_get(x_178, 0); lean_inc(x_180); lean_dec(x_178); -x_181 = lean_environment_main_module(x_180); +x_181 = l_Lean_Environment_mainModule(x_180); +lean_dec(x_180); x_182 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__3; lean_inc(x_147); x_183 = lean_alloc_ctor(2, 2, 0); @@ -41828,7 +41858,8 @@ lean_dec(x_260); x_263 = lean_ctor_get(x_261, 0); lean_inc(x_263); lean_dec(x_261); -x_264 = lean_environment_main_module(x_263); +x_264 = l_Lean_Environment_mainModule(x_263); +lean_dec(x_263); x_265 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__6; lean_inc(x_147); x_266 = lean_alloc_ctor(2, 2, 0); @@ -42516,7 +42547,8 @@ lean_dec(x_17); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_15); x_23 = lean_alloc_ctor(2, 2, 0); @@ -42725,7 +42757,8 @@ lean_dec(x_37); x_40 = lean_ctor_get(x_38, 0); lean_inc(x_40); lean_dec(x_38); -x_41 = lean_environment_main_module(x_40); +x_41 = l_Lean_Environment_mainModule(x_40); +lean_dec(x_40); x_42 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_35); x_43 = lean_alloc_ctor(2, 2, 0); @@ -43080,7 +43113,8 @@ lean_dec(x_22); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_environment_main_module(x_25); +x_26 = l_Lean_Environment_mainModule(x_25); +lean_dec(x_25); x_27 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_20); x_28 = lean_alloc_ctor(2, 2, 0); @@ -43238,7 +43272,8 @@ lean_dec(x_18); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -x_22 = lean_environment_main_module(x_21); +x_22 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_23 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_16); x_24 = lean_alloc_ctor(2, 2, 0); @@ -43340,7 +43375,8 @@ lean_dec(x_18); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -x_22 = lean_environment_main_module(x_21); +x_22 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_23 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__1; lean_inc(x_16); x_24 = lean_alloc_ctor(2, 2, 0); diff --git a/stage0/stdlib/Lean/Elab/Extra.c b/stage0/stdlib/Lean/Elab/Extra.c index 0b819e21ac6e..572e00716fcb 100644 --- a/stage0/stdlib/Lean/Elab/Extra.c +++ b/stage0/stdlib/Lean/Elab/Extra.c @@ -369,7 +369,7 @@ lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___sp static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp___spec__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange__1___closed__1; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processBinOp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1480,7 +1480,8 @@ x_31 = lean_ctor_get(x_28, 1); x_32 = lean_ctor_get(x_30, 0); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_environment_main_module(x_32); +x_33 = l_Lean_Environment_mainModule(x_32); +lean_dec(x_32); x_34 = l_Lean_Elab_Term_elabForIn___closed__6; lean_inc(x_26); lean_ctor_set_tag(x_28, 2); @@ -1545,7 +1546,8 @@ lean_dec(x_28); x_59 = lean_ctor_get(x_57, 0); lean_inc(x_59); lean_dec(x_57); -x_60 = lean_environment_main_module(x_59); +x_60 = l_Lean_Environment_mainModule(x_59); +lean_dec(x_59); x_61 = l_Lean_Elab_Term_elabForIn___closed__6; lean_inc(x_26); x_62 = lean_alloc_ctor(2, 2, 0); @@ -1629,7 +1631,8 @@ if (lean_is_exclusive(x_90)) { x_94 = lean_ctor_get(x_91, 0); lean_inc(x_94); lean_dec(x_91); -x_95 = lean_environment_main_module(x_94); +x_95 = l_Lean_Environment_mainModule(x_94); +lean_dec(x_94); x_96 = l_Lean_Elab_Term_elabForIn___closed__6; lean_inc(x_88); if (lean_is_scalar(x_93)) { @@ -4070,7 +4073,8 @@ x_31 = lean_ctor_get(x_28, 1); x_32 = lean_ctor_get(x_30, 0); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_environment_main_module(x_32); +x_33 = l_Lean_Environment_mainModule(x_32); +lean_dec(x_32); x_34 = l_Lean_Elab_Term_elabForIn___closed__6; lean_inc(x_26); lean_ctor_set_tag(x_28, 2); @@ -4135,7 +4139,8 @@ lean_dec(x_28); x_59 = lean_ctor_get(x_57, 0); lean_inc(x_59); lean_dec(x_57); -x_60 = lean_environment_main_module(x_59); +x_60 = l_Lean_Environment_mainModule(x_59); +lean_dec(x_59); x_61 = l_Lean_Elab_Term_elabForIn___closed__6; lean_inc(x_26); x_62 = lean_alloc_ctor(2, 2, 0); @@ -4219,7 +4224,8 @@ if (lean_is_exclusive(x_90)) { x_94 = lean_ctor_get(x_91, 0); lean_inc(x_94); lean_dec(x_91); -x_95 = lean_environment_main_module(x_94); +x_95 = l_Lean_Environment_mainModule(x_94); +lean_dec(x_94); x_96 = l_Lean_Elab_Term_elabForIn___closed__6; lean_inc(x_88); if (lean_is_scalar(x_93)) { diff --git a/stage0/stdlib/Lean/Elab/Inductive.c b/stage0/stdlib/Lean/Elab/Inductive.c index 3ab50f8ce07c..7e881c206479 100644 --- a/stage0/stdlib/Lean/Elab/Inductive.c +++ b/stage0/stdlib/Lean/Elab/Inductive.c @@ -302,7 +302,7 @@ lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_objec static lean_object* l___regBuiltin_Lean_Elab_Command_elabInductiveCommand__1___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_replaceArrowBinderNames(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_Elab_MutualInductive_0__Lean_Elab_Command_checkParamsAndResultType___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_inductiveSyntaxToView___spec__18___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1071,7 +1071,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -1253,7 +1254,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); diff --git a/stage0/stdlib/Lean/Elab/InfoTree/Main.c b/stage0/stdlib/Lean/Elab/InfoTree/Main.c index 800d9ec02ff3..eb4deffb6f24 100644 --- a/stage0/stdlib/Lean/Elab/InfoTree/Main.c +++ b/stage0/stdlib/Lean/Elab/InfoTree/Main.c @@ -113,7 +113,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext(lean_object*); static lean_object* l_Lean_Elab_InfoTree_format___closed__6; static lean_object* l_Lean_getConstInfo___at_Lean_Elab_realizeGlobalConstWithInfos___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_format(lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Elab_FieldInfo_format___lambda__1___closed__2; static lean_object* l_Lean_Elab_DelabTermInfo_format___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___rarg(lean_object*, lean_object*); @@ -156,13 +156,11 @@ static lean_object* l_Lean_Elab_FVarAliasInfo_format___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_CompletionInfo_stx(lean_object*); static lean_object* l_Option_format___at_Lean_Elab_CompletionInfo_format___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___rarg___lambda__1(lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_InfoTree_substitute___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatStxRange_fmtPos___closed__6; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_InfoTree_findInfo_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); extern lean_object* l_instInhabitedPUnit; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Elab_ContextInfo_ppGoals___closed__3; static lean_object* l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_instToFormatCustomInfo; @@ -213,6 +211,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_assignInfoHoleId___rarg(lean_object*, lean_ LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___rarg___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_InfoTree_substitute___spec__6(lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at_Lean_Elab_InfoTree_findInfo_x3f___spec__2___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_CompletionInfo_lctx(lean_object*); @@ -459,6 +458,7 @@ static lean_object* l_Lean_Elab_InfoTree_format___closed__4; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_TermInfo_format___lambda__2___closed__1; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_withInfoContext___spec__1(lean_object*); static lean_object* l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree(lean_object*); @@ -2414,7 +2414,7 @@ lean_dec(x_17); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = l_Lean_Kernel_isDiagnosticsEnabled(x_20); +x_21 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_20); lean_dec(x_20); if (x_21 == 0) { @@ -2468,7 +2468,7 @@ lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean x_27 = lean_ctor_get(x_24, 0); x_28 = lean_ctor_get(x_24, 4); lean_dec(x_28); -x_29 = l_Lean_Kernel_enableDiag(x_27, x_16); +x_29 = l_Lean_Kernel_Environment_enableDiag(x_27, x_16); lean_ctor_set(x_24, 4, x_5); lean_ctor_set(x_24, 0, x_29); x_30 = lean_st_ref_set(x_8, x_24, x_25); @@ -2497,7 +2497,7 @@ lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_dec(x_24); -x_41 = l_Lean_Kernel_enableDiag(x_34, x_16); +x_41 = l_Lean_Kernel_Environment_enableDiag(x_34, x_16); x_42 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_35); @@ -2579,7 +2579,7 @@ lean_dec(x_70); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = l_Lean_Kernel_isDiagnosticsEnabled(x_73); +x_74 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_73); lean_dec(x_73); if (x_74 == 0) { @@ -2654,7 +2654,7 @@ if (lean_is_exclusive(x_77)) { lean_dec_ref(x_77); x_86 = lean_box(0); } -x_87 = l_Lean_Kernel_enableDiag(x_79, x_69); +x_87 = l_Lean_Kernel_Environment_enableDiag(x_79, x_69); if (lean_is_scalar(x_86)) { x_88 = lean_alloc_ctor(0, 8, 0); } else { @@ -2925,7 +2925,7 @@ lean_dec(x_72); x_75 = lean_ctor_get(x_73, 0); lean_inc(x_75); lean_dec(x_73); -x_76 = l_Lean_Kernel_isDiagnosticsEnabled(x_75); +x_76 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_75); lean_dec(x_75); if (x_76 == 0) { @@ -3115,7 +3115,7 @@ x_82 = lean_ctor_get(x_79, 0); x_83 = lean_ctor_get(x_79, 4); lean_dec(x_83); x_84 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__15; -x_85 = l_Lean_Kernel_enableDiag(x_82, x_84); +x_85 = l_Lean_Kernel_Environment_enableDiag(x_82, x_84); lean_ctor_set(x_79, 4, x_52); lean_ctor_set(x_79, 0, x_85); x_86 = lean_st_ref_set(x_70, x_79, x_80); @@ -3230,7 +3230,7 @@ lean_inc(x_109); lean_inc(x_108); lean_dec(x_79); x_115 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__15; -x_116 = l_Lean_Kernel_enableDiag(x_108, x_115); +x_116 = l_Lean_Kernel_Environment_enableDiag(x_108, x_115); x_117 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_117, 0, x_116); lean_ctor_set(x_117, 1, x_109); @@ -9832,8 +9832,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -9873,8 +9872,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Elab/InheritDoc.c b/stage0/stdlib/Lean/Elab/InheritDoc.c index 0c8deb1316a6..139047cffd5d 100644 --- a/stage0/stdlib/Lean/Elab/InheritDoc.c +++ b/stage0/stdlib/Lean/Elab/InheritDoc.c @@ -44,7 +44,7 @@ lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_3____lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_3____spec__2(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_3____spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_3____spec__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); @@ -232,8 +232,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -273,8 +272,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Elab/LetRec.c b/stage0/stdlib/Lean/Elab/LetRec.c index 54ffd1b006c3..07062b427497 100644 --- a/stage0/stdlib/Lean/Elab/LetRec.c +++ b/stage0/stdlib/Lean/Elab/LetRec.c @@ -243,7 +243,7 @@ static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_Le 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_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange__1___closed__5; lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3683,7 +3683,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -3865,7 +3866,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); diff --git a/stage0/stdlib/Lean/Elab/Match.c b/stage0/stdlib/Lean/Elab/Match.c index 572f5a9e3e1f..a00d7c6b2255 100644 --- a/stage0/stdlib/Lean/Elab/Match.c +++ b/stage0/stdlib/Lean/Elab/Match.c @@ -225,7 +225,7 @@ static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportMatcherResultErrors___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_normLitValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_tryPostponeIfDiscrTypeIsMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_instantiatePatternMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_matchPatternAttr; @@ -327,12 +327,10 @@ lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Intro_0__Lea static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__16___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_updateMatchType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_Pattern_toExpr_visit(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___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_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goIndex___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchOptMotive___boxed(lean_object*); extern lean_object* l_instInhabitedPUnit; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabNoMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabNoMatch___closed__1; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___rarg___closed__1; @@ -497,6 +495,7 @@ lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_ReaderT_instMonadLift(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__2; +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__22(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_mapMUnsafe_map___at_Lean_Elab_Term_elabNoFun___spec__1___lambda__1___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_reportMatcherResultErrors___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*); @@ -880,7 +879,7 @@ lean_object* l_Lean_Meta_eraseInaccessibleAnnotations___lambda__2___boxed(lean_o static lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_addVar___closed__1; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__6(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_PrettyPrinter_delabCore___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__23(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__10(lean_object*, lean_object*, lean_object*); @@ -1127,6 +1126,7 @@ lean_object* l_Lean_LocalDecl_toExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ToDepElimPattern_main___spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_processVar___closed__1; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__10(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__5(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_containsFVar___spec__1(lean_object*, lean_object*, size_t, size_t); @@ -1495,7 +1495,8 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_main_module(x_8); +x_9 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_10 = lean_ctor_get(x_2, 10); lean_inc(x_10); lean_dec(x_2); @@ -1514,7 +1515,8 @@ lean_dec(x_5); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -x_15 = lean_environment_main_module(x_14); +x_15 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_16 = lean_ctor_get(x_2, 10); lean_inc(x_16); lean_dec(x_2); @@ -11597,7 +11599,8 @@ x_33 = lean_ctor_get(x_30, 1); x_34 = lean_ctor_get(x_32, 0); lean_inc(x_34); lean_dec(x_32); -x_35 = lean_environment_find(x_34, x_29); +x_35 = l_Lean_Environment_find_x3f(x_34, x_29); +lean_dec(x_29); if (lean_obj_tag(x_35) == 0) { lean_object* x_36; @@ -11912,7 +11915,8 @@ lean_dec(x_30); x_99 = lean_ctor_get(x_97, 0); lean_inc(x_99); lean_dec(x_97); -x_100 = lean_environment_find(x_99, x_29); +x_100 = l_Lean_Environment_find_x3f(x_99, x_29); +lean_dec(x_29); if (lean_obj_tag(x_100) == 0) { lean_object* x_101; lean_object* x_102; @@ -12244,7 +12248,8 @@ if (lean_is_exclusive(x_160)) { x_164 = lean_ctor_get(x_161, 0); lean_inc(x_164); lean_dec(x_161); -x_165 = lean_environment_find(x_164, x_159); +x_165 = l_Lean_Environment_find_x3f(x_164, x_159); +lean_dec(x_159); if (lean_obj_tag(x_165) == 0) { lean_object* x_166; lean_object* x_167; @@ -12709,7 +12714,8 @@ if (lean_is_exclusive(x_243)) { x_247 = lean_ctor_get(x_244, 0); lean_inc(x_247); lean_dec(x_244); -x_248 = lean_environment_find(x_247, x_242); +x_248 = l_Lean_Environment_find_x3f(x_247, x_242); +lean_dec(x_242); if (lean_obj_tag(x_248) == 0) { lean_object* x_249; lean_object* x_250; @@ -14798,7 +14804,8 @@ x_30 = lean_ctor_get(x_27, 1); x_31 = lean_ctor_get(x_29, 0); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_environment_find(x_31, x_26); +x_32 = l_Lean_Environment_find_x3f(x_31, x_26); +lean_dec(x_26); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; @@ -15033,7 +15040,8 @@ lean_dec(x_27); x_79 = lean_ctor_get(x_77, 0); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_find(x_79, x_26); +x_80 = l_Lean_Environment_find_x3f(x_79, x_26); +lean_dec(x_26); if (lean_obj_tag(x_80) == 0) { lean_object* x_81; lean_object* x_82; @@ -15288,7 +15296,8 @@ if (lean_is_exclusive(x_126)) { x_130 = lean_ctor_get(x_127, 0); lean_inc(x_130); lean_dec(x_127); -x_131 = lean_environment_find(x_130, x_125); +x_131 = l_Lean_Environment_find_x3f(x_130, x_125); +lean_dec(x_125); if (lean_obj_tag(x_131) == 0) { lean_object* x_132; lean_object* x_133; @@ -15815,7 +15824,8 @@ x_16 = lean_ctor_get(x_13, 1); x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -x_18 = lean_environment_find(x_17, x_12); +x_18 = l_Lean_Environment_find_x3f(x_17, x_12); +lean_dec(x_12); if (lean_obj_tag(x_18) == 0) { lean_dec(x_11); @@ -16254,7 +16264,8 @@ lean_dec(x_13); x_114 = lean_ctor_get(x_112, 0); lean_inc(x_114); lean_dec(x_112); -x_115 = lean_environment_find(x_114, x_12); +x_115 = l_Lean_Environment_find_x3f(x_114, x_12); +lean_dec(x_12); if (lean_obj_tag(x_115) == 0) { lean_object* x_116; @@ -16533,7 +16544,8 @@ if (lean_is_exclusive(x_168)) { x_172 = lean_ctor_get(x_169, 0); lean_inc(x_172); lean_dec(x_169); -x_173 = lean_environment_find(x_172, x_167); +x_173 = l_Lean_Environment_find_x3f(x_172, x_167); +lean_dec(x_167); if (lean_obj_tag(x_173) == 0) { lean_object* x_174; @@ -23130,7 +23142,8 @@ lean_dec(x_28); x_31 = lean_ctor_get(x_29, 0); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_environment_find(x_31, x_27); +x_32 = l_Lean_Environment_find_x3f(x_31, x_27); +lean_dec(x_27); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; @@ -24535,7 +24548,8 @@ lean_dec(x_20); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_find(x_23, x_18); +x_24 = l_Lean_Environment_find_x3f(x_23, x_18); +lean_dec(x_18); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; @@ -24725,7 +24739,8 @@ lean_dec(x_77); x_80 = lean_ctor_get(x_78, 0); lean_inc(x_80); lean_dec(x_78); -x_81 = lean_environment_find(x_80, x_75); +x_81 = l_Lean_Environment_find_x3f(x_80, x_75); +lean_dec(x_75); if (lean_obj_tag(x_81) == 0) { lean_object* x_82; lean_object* x_83; @@ -29793,7 +29808,7 @@ lean_dec(x_15); x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); -x_19 = l_Lean_Kernel_isDiagnosticsEnabled(x_18); +x_19 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_18); lean_dec(x_18); if (x_19 == 0) { @@ -29843,7 +29858,7 @@ lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean x_25 = lean_ctor_get(x_22, 0); x_26 = lean_ctor_get(x_22, 4); lean_dec(x_26); -x_27 = l_Lean_Kernel_enableDiag(x_25, x_14); +x_27 = l_Lean_Kernel_Environment_enableDiag(x_25, x_14); x_28 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__6; lean_ctor_set(x_22, 4, x_28); lean_ctor_set(x_22, 0, x_27); @@ -29873,7 +29888,7 @@ lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); lean_dec(x_22); -x_40 = l_Lean_Kernel_enableDiag(x_33, x_14); +x_40 = l_Lean_Kernel_Environment_enableDiag(x_33, x_14); x_41 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__6; x_42 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_42, 0, x_40); @@ -45204,7 +45219,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -45386,7 +45402,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); @@ -48367,7 +48384,8 @@ x_20 = lean_ctor_get(x_17, 1); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -x_22 = lean_environment_main_module(x_21); +x_22 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__2___closed__1; lean_inc(x_15); lean_ctor_set_tag(x_17, 2); @@ -48411,7 +48429,8 @@ x_39 = lean_ctor_get(x_37, 0); x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); lean_dec(x_39); -x_41 = lean_environment_main_module(x_40); +x_41 = l_Lean_Environment_mainModule(x_40); +lean_dec(x_40); x_42 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__4; lean_inc(x_15); x_43 = lean_alloc_ctor(2, 2, 0); @@ -48450,7 +48469,8 @@ lean_dec(x_37); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__4; lean_inc(x_15); x_57 = lean_alloc_ctor(2, 2, 0); @@ -48520,7 +48540,8 @@ lean_dec(x_17); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__2___closed__1; lean_inc(x_15); x_76 = lean_alloc_ctor(2, 2, 0); @@ -48571,7 +48592,8 @@ if (lean_is_exclusive(x_90)) { x_94 = lean_ctor_get(x_91, 0); lean_inc(x_94); lean_dec(x_91); -x_95 = lean_environment_main_module(x_94); +x_95 = l_Lean_Environment_mainModule(x_94); +lean_dec(x_94); x_96 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__4; lean_inc(x_15); x_97 = lean_alloc_ctor(2, 2, 0); @@ -50188,8 +50210,7 @@ x_27 = lean_ctor_get(x_24, 1); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -lean_inc(x_23); -x_29 = lean_environment_find(x_28, x_23); +x_29 = l_Lean_Environment_find_x3f(x_28, x_23); if (lean_obj_tag(x_29) == 0) { uint8_t x_30; lean_object* x_31; @@ -50302,8 +50323,7 @@ lean_dec(x_24); x_55 = lean_ctor_get(x_53, 0); lean_inc(x_55); lean_dec(x_53); -lean_inc(x_23); -x_56 = lean_environment_find(x_55, x_23); +x_56 = l_Lean_Environment_find_x3f(x_55, x_23); if (lean_obj_tag(x_56) == 0) { uint8_t x_57; lean_object* x_58; lean_object* x_59; @@ -50698,7 +50718,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -50880,7 +50901,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); @@ -52055,7 +52077,8 @@ x_19 = lean_ctor_get(x_16, 1); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__2___closed__1; lean_inc(x_14); lean_ctor_set_tag(x_16, 2); @@ -52097,7 +52120,8 @@ x_36 = lean_ctor_get(x_34, 0); x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); lean_dec(x_36); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__4; lean_inc(x_14); x_40 = lean_alloc_ctor(2, 2, 0); @@ -52136,7 +52160,8 @@ lean_dec(x_34); x_51 = lean_ctor_get(x_49, 0); lean_inc(x_51); lean_dec(x_49); -x_52 = lean_environment_main_module(x_51); +x_52 = l_Lean_Environment_mainModule(x_51); +lean_dec(x_51); x_53 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__4; lean_inc(x_14); x_54 = lean_alloc_ctor(2, 2, 0); @@ -52206,7 +52231,8 @@ lean_dec(x_16); x_70 = lean_ctor_get(x_68, 0); lean_inc(x_70); lean_dec(x_68); -x_71 = lean_environment_main_module(x_70); +x_71 = l_Lean_Environment_mainModule(x_70); +lean_dec(x_70); x_72 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__2___closed__1; lean_inc(x_14); x_73 = lean_alloc_ctor(2, 2, 0); @@ -52255,7 +52281,8 @@ if (lean_is_exclusive(x_85)) { x_89 = lean_ctor_get(x_86, 0); lean_inc(x_89); lean_dec(x_86); -x_90 = lean_environment_main_module(x_89); +x_90 = l_Lean_Environment_mainModule(x_89); +lean_dec(x_89); x_91 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__4; lean_inc(x_14); x_92 = lean_alloc_ctor(2, 2, 0); @@ -53656,7 +53683,8 @@ x_10 = lean_ctor_get(x_8, 0); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); -x_12 = lean_environment_main_module(x_11); +x_12 = l_Lean_Environment_mainModule(x_11); +lean_dec(x_11); x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabNoFun___spec__1___lambda__1___closed__3; x_14 = l_Lean_addMacroScope(x_12, x_13, x_7); x_15 = lean_box(0); @@ -53680,7 +53708,8 @@ lean_dec(x_8); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabNoFun___spec__1___lambda__1___closed__3; x_23 = l_Lean_addMacroScope(x_21, x_22, x_7); x_24 = lean_box(0); diff --git a/stage0/stdlib/Lean/Elab/Open.c b/stage0/stdlib/Lean/Elab/Open.c index 3defbf6367a9..3214c1fab7c6 100644 --- a/stage0/stdlib/Lean/Elab/Open.c +++ b/stage0/stdlib/Lean/Elab/Open.c @@ -136,7 +136,7 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_OpenDecl_elabOpenDec LEAN_EXPORT lean_object* l_Lean_resolveNamespaceCore___at_Lean_Elab_OpenDecl_elabOpenDecl___spec__39___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveUniqueNamespace___at_Lean_Elab_OpenDecl_elabOpenDecl___spec__6___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_OpenDecl_elabOpenDecl___spec__41(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_resolveNameUsingNamespaces(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_addOpenDecl___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_OpenDecl_elabOpenDecl___spec__37___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3896,8 +3896,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_OpenDecl_elabOpenDec _start: { lean_object* x_8; -lean_inc(x_1); -x_8 = lean_environment_find(x_7, x_1); +x_8 = l_Lean_Environment_find_x3f(x_7, x_1); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; @@ -4268,8 +4267,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_OpenDecl_elabOpenDec _start: { lean_object* x_8; -lean_inc(x_1); -x_8 = lean_environment_find(x_7, x_1); +x_8 = l_Lean_Environment_find_x3f(x_7, x_1); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; @@ -5699,8 +5697,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_OpenDecl_elabOpenDec _start: { lean_object* x_8; -lean_inc(x_1); -x_8 = lean_environment_find(x_7, x_1); +x_8 = l_Lean_Environment_find_x3f(x_7, x_1); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; @@ -6714,8 +6711,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_OpenDecl_elabOpenDec _start: { lean_object* x_8; -lean_inc(x_1); -x_8 = lean_environment_find(x_7, x_1); +x_8 = l_Lean_Environment_find_x3f(x_7, x_1); if (lean_obj_tag(x_8) == 0) { uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; diff --git a/stage0/stdlib/Lean/Elab/PatternVar.c b/stage0/stdlib/Lean/Elab/PatternVar.c index cbc745999a88..f9f9906a268c 100644 --- a/stage0/stdlib/Lean/Elab/PatternVar.c +++ b/stage0/stdlib/Lean/Elab/PatternVar.c @@ -109,7 +109,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_PatternVar_0 LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected_showName___spec__3___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected_showName___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppCore___spec__4(lean_object*); extern lean_object* l_Lean_matchPatternAttr; lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); @@ -360,7 +360,7 @@ lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_objec lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -393,7 +393,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_pushNewArg_ static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg___lambda__4___closed__4; uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__25; @@ -467,6 +466,7 @@ static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPa static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__10; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg___closed__2; +lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_isDone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3184,10 +3184,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_Coll _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_3 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); lean_dec(x_1); -x_4 = l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(x_3, x_2); +x_4 = l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(x_3, x_2); x_5 = 1; x_6 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected_showName___closed__3; lean_inc(x_2); @@ -4962,7 +4962,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_Coll _start: { uint8_t x_13; -lean_inc(x_2); lean_inc(x_1); x_13 = l_Lean_Environment_isConstructor(x_1, x_2); if (x_13 == 0) @@ -5298,7 +5297,7 @@ x_16 = lean_ctor_get(x_13, 0); lean_inc(x_16); lean_dec(x_13); x_17 = lean_box(0); -x_18 = lean_ctor_get(x_16, 1); +x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_inc(x_16); x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg___lambda__3___boxed), 13, 2); @@ -8324,8 +8323,7 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_1); -x_15 = lean_environment_find(x_14, x_1); +x_15 = l_Lean_Environment_find_x3f(x_14, x_1); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -8365,8 +8363,7 @@ lean_dec(x_10); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_1); -x_27 = lean_environment_find(x_26, x_1); +x_27 = l_Lean_Environment_find_x3f(x_26, x_1); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; @@ -12159,7 +12156,8 @@ x_26 = lean_ctor_get(x_24, 0); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); lean_dec(x_26); -x_28 = lean_environment_main_module(x_27); +x_28 = l_Lean_Environment_mainModule(x_27); +lean_dec(x_27); x_29 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__5; x_30 = l_Lean_addMacroScope(x_28, x_29, x_23); x_31 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__2; @@ -12188,7 +12186,8 @@ lean_dec(x_24); x_39 = lean_ctor_get(x_37, 0); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); +x_40 = l_Lean_Environment_mainModule(x_39); +lean_dec(x_39); x_41 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__5; x_42 = l_Lean_addMacroScope(x_40, x_41, x_23); x_43 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__2; @@ -12344,7 +12343,8 @@ lean_dec(x_24); x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); -x_28 = lean_environment_main_module(x_27); +x_28 = l_Lean_Environment_mainModule(x_27); +lean_dec(x_27); x_29 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___closed__3; x_30 = l_Lean_addMacroScope(x_28, x_29, x_23); x_31 = lean_box(0); @@ -14121,8 +14121,7 @@ lean_dec(x_20); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -lean_inc(x_19); -x_24 = lean_environment_find(x_23, x_19); +x_24 = l_Lean_Environment_find_x3f(x_23, x_19); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; @@ -14227,8 +14226,7 @@ lean_dec(x_41); x_44 = lean_ctor_get(x_42, 0); lean_inc(x_44); lean_dec(x_42); -lean_inc(x_40); -x_45 = lean_environment_find(x_44, x_40); +x_45 = l_Lean_Environment_find_x3f(x_44, x_40); if (lean_obj_tag(x_45) == 0) { lean_object* x_46; lean_object* x_47; @@ -16550,7 +16548,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_13); +x_31 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_32 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_32, 0, x_25); lean_ctor_set(x_32, 1, x_31); @@ -16728,7 +16727,8 @@ lean_dec(x_26); x_80 = lean_ctor_get(x_78, 1); lean_inc(x_80); lean_dec(x_78); -x_81 = lean_environment_main_module(x_13); +x_81 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_82 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_82, 0, x_25); lean_ctor_set(x_82, 1, x_81); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Basic.c b/stage0/stdlib/Lean/Elab/PreDefinition/Basic.c index 25f7019e962e..cae363382202 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Basic.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Basic.c @@ -112,12 +112,10 @@ lean_object* l_Lean_CollectLevelParams_main(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_checkCodomainsLevel___spec__2___lambda__1___closed__2; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_shareCommonPreDefs___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_checkCodomainsLevel___spec__2___lambda__2___closed__2; -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); extern lean_object* l_instInhabitedNat; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_addAndCompileUnsafe___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addAndCompileUnsafe(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_shareCommonPreDefs___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_checkCodomainsLevel___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -164,6 +162,7 @@ static lean_object* l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_repor 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_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_diagnostics_threshold_proofSize; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_addAndCompilePartialRec___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -338,6 +337,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_fixLevelParams___spec__3___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_shareCommonPreDefs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_checkCodomainsLevel___spec__2___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_shouldGenCodeFor(lean_object*); @@ -11095,7 +11095,7 @@ lean_dec(x_28); x_31 = lean_ctor_get(x_29, 0); lean_inc(x_31); lean_dec(x_29); -x_32 = l_Lean_Kernel_isDiagnosticsEnabled(x_31); +x_32 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_31); lean_dec(x_31); if (x_32 == 0) { @@ -11141,7 +11141,7 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean x_38 = lean_ctor_get(x_35, 0); x_39 = lean_ctor_get(x_35, 4); lean_dec(x_39); -x_40 = l_Lean_Kernel_enableDiag(x_38, x_27); +x_40 = l_Lean_Kernel_Environment_enableDiag(x_38, x_27); x_41 = l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux___lambda__4___closed__4; lean_ctor_set(x_35, 4, x_41); lean_ctor_set(x_35, 0, x_40); @@ -11171,7 +11171,7 @@ lean_inc(x_48); lean_inc(x_47); lean_inc(x_46); lean_dec(x_35); -x_53 = l_Lean_Kernel_enableDiag(x_46, x_27); +x_53 = l_Lean_Kernel_Environment_enableDiag(x_46, x_27); x_54 = l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux___lambda__4___closed__4; x_55 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_55, 0, x_53); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/EqUnfold.c b/stage0/stdlib/Lean/Elab/PreDefinition/EqUnfold.c index 64f3468c4d74..9a623ebaa9b5 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/EqUnfold.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/EqUnfold.c @@ -48,9 +48,7 @@ extern lean_object* l_Lean_maxRecDepth; static lean_object* l_Lean_Meta_initFn____x40_Lean_Elab_PreDefinition_EqUnfold___hyg_669____lambda__2___closed__8; uint8_t lean_string_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); extern lean_object* l_Lean_Meta_smartUnfolding; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_getConstUnfoldEqnFor_x3f___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_getConstUnfoldEqnFor_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getConstUnfoldEqnFor_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -63,6 +61,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getConstUnfoldEqnFor_x3f___lambda__3(lean_o static lean_object* l_Lean_Meta_getConstUnfoldEqnFor_x3f___lambda__2___closed__2; lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Meta_getConstUnfoldEqnFor_x3f___lambda__2___closed__3; static lean_object* l_Lean_Meta_initFn____x40_Lean_Elab_PreDefinition_EqUnfold___hyg_669____lambda__2___closed__5; lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -118,6 +117,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Elab_PreDefinition_EqUnf lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Meta_getConstUnfoldEqnFor_x3f___lambda__2___closed__1; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Meta_initFn____x40_Lean_Elab_PreDefinition_EqUnfold___hyg_669____lambda__2___closed__9; uint8_t lean_usize_dec_lt(size_t, size_t); 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*); @@ -456,7 +456,7 @@ lean_dec(x_13); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -x_17 = l_Lean_Kernel_isDiagnosticsEnabled(x_16); +x_17 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_16); lean_dec(x_16); if (x_17 == 0) { @@ -506,7 +506,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean x_23 = lean_ctor_get(x_20, 0); x_24 = lean_ctor_get(x_20, 4); lean_dec(x_24); -x_25 = l_Lean_Kernel_enableDiag(x_23, x_12); +x_25 = l_Lean_Kernel_Environment_enableDiag(x_23, x_12); x_26 = l_Lean_Meta_tryURefl___closed__5; lean_ctor_set(x_20, 4, x_26); lean_ctor_set(x_20, 0, x_25); @@ -536,7 +536,7 @@ lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); lean_dec(x_20); -x_38 = l_Lean_Kernel_enableDiag(x_31, x_12); +x_38 = l_Lean_Kernel_Environment_enableDiag(x_31, x_12); x_39 = l_Lean_Meta_tryURefl___closed__5; x_40 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_40, 0, x_38); @@ -2222,7 +2222,6 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_5); x_12 = l_Lean_Environment_isSafeDefinition(x_11, x_5); if (x_12 == 0) { @@ -2257,7 +2256,6 @@ lean_dec(x_7); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -lean_inc(x_5); x_20 = l_Lean_Environment_isSafeDefinition(x_19, x_5); if (x_20 == 0) { diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c index 852d70ebe301..7462e1d67384 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c @@ -162,14 +162,12 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreD LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__28___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_backward_eqns_deepRecursiveSplit; static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6256____closed__14; -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldrM___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__7; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_smartUnfolding; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); lean_object* l_Lean_Expr_replaceFVarId(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -238,6 +236,7 @@ LEAN_EXPORT lean_object* l_Lean_LocalContext_foldrM___at___private_Lean_Elab_Pre LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldrMAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2(lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -518,6 +517,7 @@ uint8_t l_Lean_isAuxRecursorWithSuffix(lean_object*, lean_object*, lean_object*) lean_object* lean_array_get_size(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__5___closed__4; static lean_object* l_Lean_Elab_Eqns_expandRHS_x3f___closed__2; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__2; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3624,7 +3624,7 @@ lean_dec(x_13); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -x_17 = l_Lean_Kernel_isDiagnosticsEnabled(x_16); +x_17 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_16); lean_dec(x_16); if (x_17 == 0) { @@ -3674,7 +3674,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean x_23 = lean_ctor_get(x_20, 0); x_24 = lean_ctor_get(x_20, 4); lean_dec(x_24); -x_25 = l_Lean_Kernel_enableDiag(x_23, x_12); +x_25 = l_Lean_Kernel_Environment_enableDiag(x_23, x_12); x_26 = l_Lean_Elab_Eqns_tryURefl___closed__5; lean_ctor_set(x_20, 4, x_26); lean_ctor_set(x_20, 0, x_25); @@ -3704,7 +3704,7 @@ lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); lean_dec(x_20); -x_38 = l_Lean_Kernel_enableDiag(x_31, x_12); +x_38 = l_Lean_Kernel_Environment_enableDiag(x_31, x_12); x_39 = l_Lean_Elab_Eqns_tryURefl___closed__5; x_40 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_40, 0, x_38); @@ -30761,7 +30761,7 @@ lean_dec(x_15); x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); -x_19 = l_Lean_Kernel_isDiagnosticsEnabled(x_18); +x_19 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_18); lean_dec(x_18); if (x_19 == 0) { @@ -30812,7 +30812,7 @@ x_25 = lean_ctor_get(x_21, 1); x_26 = lean_ctor_get(x_23, 0); x_27 = lean_ctor_get(x_23, 4); lean_dec(x_27); -x_28 = l_Lean_Kernel_enableDiag(x_26, x_14); +x_28 = l_Lean_Kernel_Environment_enableDiag(x_26, x_14); lean_inc(x_3); lean_ctor_set(x_21, 1, x_3); lean_ctor_set(x_21, 0, x_3); @@ -30845,7 +30845,7 @@ lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_dec(x_23); -x_41 = l_Lean_Kernel_enableDiag(x_34, x_14); +x_41 = l_Lean_Kernel_Environment_enableDiag(x_34, x_14); lean_inc(x_3); lean_ctor_set(x_21, 1, x_3); lean_ctor_set(x_21, 0, x_3); @@ -30903,7 +30903,7 @@ if (lean_is_exclusive(x_47)) { lean_dec_ref(x_47); x_56 = lean_box(0); } -x_57 = l_Lean_Kernel_enableDiag(x_49, x_14); +x_57 = l_Lean_Kernel_Environment_enableDiag(x_49, x_14); lean_inc(x_3); x_58 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_58, 0, x_3); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Nonrec/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Nonrec/Eqns.c index 38e05f8051ae..595d466a028f 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Nonrec/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Nonrec/Eqns.c @@ -32,7 +32,7 @@ uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Nonrec_mkEqns___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Nonrec_Eqns_0__Lean_Elab_Nonrec_mkProof_go___lambda__1___closed__17; lean_object* lean_array_fget(lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_tactic_hygienic; static lean_object* l___private_Lean_Elab_PreDefinition_Nonrec_Eqns_0__Lean_Elab_Nonrec_mkProof_go___lambda__1___closed__4; lean_object* l_Lean_Elab_Eqns_tryContradiction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -43,10 +43,8 @@ extern lean_object* l_Lean_maxRecDepth; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Nonrec_Eqns_0__Lean_Elab_Nonrec_mkProof___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Nonrec_mkEqns___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Elab_Nonrec_mkEqns___lambda__1___closed__1; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Elab_Nonrec_mkEqns___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Nonrec_getEqnsFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Nonrec_mkEqns___closed__1; @@ -63,6 +61,7 @@ lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(le static lean_object* l___private_Lean_Elab_PreDefinition_Nonrec_Eqns_0__Lean_Elab_Nonrec_mkProof_go___lambda__1___closed__14; lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Nonrec_Eqns_0__Lean_Elab_Nonrec_mkSimpleEqThm___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_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Nonrec_getEqnsFor_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -145,6 +144,7 @@ lean_object* l_Lean_Elab_Eqns_tryURefl(lean_object*, lean_object*, lean_object*, lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); uint8_t lean_nat_dec_le(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*); uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); @@ -699,8 +699,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; @@ -764,8 +763,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; @@ -4862,7 +4860,7 @@ lean_dec(x_19); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -x_23 = l_Lean_Kernel_isDiagnosticsEnabled(x_22); +x_23 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_22); lean_dec(x_22); if (x_23 == 0) { @@ -4908,7 +4906,7 @@ lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean x_29 = lean_ctor_get(x_26, 0); x_30 = lean_ctor_get(x_26, 4); lean_dec(x_30); -x_31 = l_Lean_Kernel_enableDiag(x_29, x_18); +x_31 = l_Lean_Kernel_Environment_enableDiag(x_29, x_18); x_32 = l_Lean_Elab_Nonrec_mkEqns___closed__3; lean_ctor_set(x_26, 4, x_32); lean_ctor_set(x_26, 0, x_31); @@ -4938,7 +4936,7 @@ lean_inc(x_39); lean_inc(x_38); lean_inc(x_37); lean_dec(x_26); -x_44 = l_Lean_Kernel_enableDiag(x_37, x_18); +x_44 = l_Lean_Kernel_Environment_enableDiag(x_37, x_18); x_45 = l_Lean_Elab_Nonrec_mkEqns___closed__3; x_46 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_46, 0, x_44); @@ -5041,8 +5039,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; @@ -5482,8 +5479,7 @@ lean_dec(x_8); x_103 = lean_ctor_get(x_101, 0); lean_inc(x_103); lean_dec(x_101); -lean_inc(x_1); -x_104 = lean_environment_find(x_103, x_1); +x_104 = l_Lean_Environment_find_x3f(x_103, x_1); if (lean_obj_tag(x_104) == 0) { lean_object* x_105; lean_object* x_106; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/BRecOn.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/BRecOn.c index ce2a3e9fb396..d68f864700f3 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/BRecOn.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/BRecOn.c @@ -110,7 +110,7 @@ LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_PreDefinition_Structur LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_withBelowDict___spec__8___rarg___lambda__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_toBelowAux___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_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__35(lean_object*); static lean_object* l_Lean_Elab_Structural_searchPProd___rarg___closed__6; @@ -4343,8 +4343,7 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_2); -x_15 = lean_environment_find(x_14, x_2); +x_15 = l_Lean_Environment_find_x3f(x_14, x_2); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -4384,8 +4383,7 @@ lean_dec(x_10); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_2); -x_27 = lean_environment_find(x_26, x_2); +x_27 = l_Lean_Environment_find_x3f(x_26, x_2); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; @@ -4477,8 +4475,7 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_2); -x_15 = lean_environment_find(x_14, x_2); +x_15 = l_Lean_Environment_find_x3f(x_14, x_2); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -4518,8 +4515,7 @@ lean_dec(x_10); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_2); -x_27 = lean_environment_find(x_26, x_2); +x_27 = l_Lean_Environment_find_x3f(x_26, x_2); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c index 4aefd6591708..0cb76dc951a1 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c @@ -59,10 +59,8 @@ static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Structural_mkEqns___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* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_registerEqnsInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__19; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Structural_mkEqns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_eqnInfoExt; @@ -85,6 +83,7 @@ static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__4; lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__4; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__8; @@ -164,6 +163,7 @@ lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___s lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3366,7 +3366,7 @@ lean_dec(x_19); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -x_23 = l_Lean_Kernel_isDiagnosticsEnabled(x_22); +x_23 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_22); lean_dec(x_22); if (x_23 == 0) { @@ -3412,7 +3412,7 @@ lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean x_29 = lean_ctor_get(x_26, 0); x_30 = lean_ctor_get(x_26, 4); lean_dec(x_30); -x_31 = l_Lean_Kernel_enableDiag(x_29, x_18); +x_31 = l_Lean_Kernel_Environment_enableDiag(x_29, x_18); x_32 = l_Lean_Elab_Structural_mkEqns___closed__3; lean_ctor_set(x_26, 4, x_32); lean_ctor_set(x_26, 0, x_31); @@ -3442,7 +3442,7 @@ lean_inc(x_39); lean_inc(x_38); lean_inc(x_37); lean_dec(x_26); -x_44 = l_Lean_Kernel_enableDiag(x_37, x_18); +x_44 = l_Lean_Kernel_Environment_enableDiag(x_37, x_18); x_45 = l_Lean_Elab_Structural_mkEqns___closed__3; x_46 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_46, 0, x_44); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/FindRecArg.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/FindRecArg.c index 2d26fc1593cd..d9ec927a0ea4 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/FindRecArg.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/FindRecArg.c @@ -93,7 +93,7 @@ LEAN_EXPORT lean_object* l_Lean_hasConst___at_Lean_Elab_Structural_tryAllArgs___ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_FindRecArg_0__Lean_Elab_Structural_dedup___rarg___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Structural_tryAllArgs___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_getRecArgInfos___lambda__4___closed__3; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Structural_tryAllArgs___spec__9___rarg___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_tryAllArgs___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5559,7 +5559,8 @@ if (lean_is_exclusive(x_19)) { x_23 = lean_ctor_get(x_20, 0); lean_inc(x_23); lean_dec(x_20); -x_24 = lean_environment_find(x_23, x_17); +x_24 = l_Lean_Environment_find_x3f(x_23, x_17); +lean_dec(x_17); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; @@ -5615,7 +5616,6 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_30); x_37 = l_Lean_Meta_isInductivePredicate(x_30, x_7, x_8, x_9, x_10, x_35); if (lean_obj_tag(x_37) == 0) { diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c index 2f68cc9553d9..fe472498bbe9 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c @@ -75,7 +75,7 @@ extern lean_object* l_Lean_casesOnSuffix; lean_object* l_Lean_Meta_Match_MatcherInfo_arity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkIndPredBRecOn___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_List_forIn_x27_loop___at___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop_addBelow___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_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApp___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_addTrace___at___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApp___spec__3___closed__1; @@ -8658,8 +8658,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -8699,8 +8698,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Main.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Main.c index 9d7e2a979781..1b0f3771d2e9 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Main.c @@ -8840,7 +8840,6 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_21); x_22 = l_Lean_Meta_isInductivePredicate(x_21, x_5, x_6, x_7, x_8, x_19); if (lean_obj_tag(x_22) == 0) { diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c index 02b95b626854..2560faf7eb4d 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c @@ -81,12 +81,10 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_ static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_1563____closed__1; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_WF_eqnInfoExt; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_registerEqnsInfo___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__17; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_registerEqnsInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -109,6 +107,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_getUnfoldFor_x3f(lean_object*, lean_obje lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__6; lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_getUnfoldFor_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -215,6 +214,7 @@ lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___s static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_1840____closed__16; lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static uint64_t l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__20; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -4500,7 +4500,7 @@ lean_dec(x_20); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = l_Lean_Kernel_isDiagnosticsEnabled(x_23); +x_24 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_23); lean_dec(x_23); if (x_24 == 0) { @@ -4546,7 +4546,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean x_30 = lean_ctor_get(x_27, 0); x_31 = lean_ctor_get(x_27, 4); lean_dec(x_31); -x_32 = l_Lean_Kernel_enableDiag(x_30, x_19); +x_32 = l_Lean_Kernel_Environment_enableDiag(x_30, x_19); x_33 = l_Lean_Elab_WF_mkEqns___closed__3; lean_ctor_set(x_27, 4, x_33); lean_ctor_set(x_27, 0, x_32); @@ -4577,7 +4577,7 @@ lean_inc(x_40); lean_inc(x_39); lean_inc(x_38); lean_dec(x_27); -x_45 = l_Lean_Kernel_enableDiag(x_38, x_19); +x_45 = l_Lean_Kernel_Environment_enableDiag(x_38, x_19); x_46 = l_Lean_Elab_WF_mkEqns___closed__3; x_47 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_47, 0, x_45); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Fix.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Fix.c index c29c57355401..bdd2e5e989b7 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Fix.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Fix.c @@ -85,7 +85,7 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Fix_0__Lean_Elab_WF_a lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Fix_0__Lean_Elab_WF_processPSigmaCasesOn___closed__1; static lean_object* l_Lean_Elab_WF_mkFix___lambda__1___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_WF_Fix_0__Lean_Elab_WF_replaceRecApps_processApp___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_WF_applyCleanWfTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1201,8 +1201,7 @@ x_14 = lean_ctor_get(x_11, 1); x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); lean_dec(x_13); -lean_inc(x_2); -x_16 = lean_environment_find(x_15, x_2); +x_16 = l_Lean_Environment_find_x3f(x_15, x_2); if (lean_obj_tag(x_16) == 0) { uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; @@ -1242,8 +1241,7 @@ lean_dec(x_11); x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); -lean_inc(x_2); -x_28 = lean_environment_find(x_27, x_2); +x_28 = l_Lean_Environment_find_x3f(x_27, x_2); if (lean_obj_tag(x_28) == 0) { uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -1335,8 +1333,7 @@ x_14 = lean_ctor_get(x_11, 1); x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); lean_dec(x_13); -lean_inc(x_2); -x_16 = lean_environment_find(x_15, x_2); +x_16 = l_Lean_Environment_find_x3f(x_15, x_2); if (lean_obj_tag(x_16) == 0) { uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; @@ -1376,8 +1373,7 @@ lean_dec(x_11); x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); -lean_inc(x_2); -x_28 = lean_environment_find(x_27, x_2); +x_28 = l_Lean_Environment_find_x3f(x_27, x_2); if (lean_obj_tag(x_28) == 0) { uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c index dd4b36749a4a..c41e2473b207 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c @@ -196,7 +196,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_GuessLex_withRecApp LEAN_EXPORT lean_object* l_Lean_Expr_hasAnyFVar_visit___at_Lean_Elab_WF_GuessLex_complexMeasures___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_RecCallWithContext_create(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__32(lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_guessLex___lambda__1(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_WF_applyCleanWfTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_refl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4116,8 +4116,7 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_2); -x_15 = lean_environment_find(x_14, x_2); +x_15 = l_Lean_Environment_find_x3f(x_14, x_2); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -4157,8 +4156,7 @@ lean_dec(x_10); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_2); -x_27 = lean_environment_find(x_26, x_2); +x_27 = l_Lean_Environment_find_x3f(x_26, x_2); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; @@ -4258,8 +4256,7 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_2); -x_15 = lean_environment_find(x_14, x_2); +x_15 = l_Lean_Environment_find_x3f(x_14, x_2); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -4299,8 +4296,7 @@ lean_dec(x_10); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_2); -x_27 = lean_environment_find(x_26, x_2); +x_27 = l_Lean_Environment_find_x3f(x_26, x_2); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c index d089ef7833a1..dc9c62179472 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c @@ -100,9 +100,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_visitLet_visit___at_Lean_Elab_getFixedPrefi LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27___at_Lean_Elab_getFixedPrefix___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_iteToDIte___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_getFixedPrefix___spec__13(lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___lambda__2___closed__5; uint8_t l_Lean_Expr_isLambda(lean_object*); static lean_object* l_Lean_Elab_wfRecursion___lambda__8___closed__1; @@ -148,6 +146,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__9(lean_object*, size_ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3___closed__1; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___lambda__2___closed__1; lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); @@ -327,6 +326,7 @@ static lean_object* l_Lean_Elab_varyingVarNames___closed__6; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__10(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); @@ -8879,7 +8879,7 @@ lean_dec(x_60); x_63 = lean_ctor_get(x_61, 0); lean_inc(x_63); lean_dec(x_61); -x_64 = l_Lean_Kernel_isDiagnosticsEnabled(x_63); +x_64 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_63); lean_dec(x_63); if (x_64 == 0) { @@ -9122,7 +9122,7 @@ lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean x_70 = lean_ctor_get(x_67, 0); x_71 = lean_ctor_get(x_67, 4); lean_dec(x_71); -x_72 = l_Lean_Kernel_enableDiag(x_70, x_59); +x_72 = l_Lean_Kernel_Environment_enableDiag(x_70, x_59); x_73 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__4___closed__3; lean_ctor_set(x_67, 4, x_73); lean_ctor_set(x_67, 0, x_72); @@ -9202,7 +9202,7 @@ lean_inc(x_85); lean_inc(x_84); lean_inc(x_83); lean_dec(x_67); -x_90 = l_Lean_Kernel_enableDiag(x_83, x_59); +x_90 = l_Lean_Kernel_Environment_enableDiag(x_83, x_59); x_91 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__4___closed__3; x_92 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_92, 0, x_90); diff --git a/stage0/stdlib/Lean/Elab/Print.c b/stage0/stdlib/Lean/Elab/Print.c index 1a2a7d00e152..ec048768a6bf 100644 --- a/stage0/stdlib/Lean/Elab/Print.c +++ b/stage0/stdlib/Lean/Elab/Print.c @@ -90,7 +90,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrintEqns_declRange__1_ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printStructure___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__6; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printEqnsOf___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printInduct___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printStructure___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1546,8 +1546,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -1588,8 +1587,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -8282,7 +8280,6 @@ x_15 = lean_ctor_get(x_12, 1); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -lean_inc(x_1); x_17 = l_Lean_getStructureCtor(x_16, x_1); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); @@ -8443,7 +8440,6 @@ lean_dec(x_12); x_83 = lean_ctor_get(x_81, 0); lean_inc(x_83); lean_dec(x_81); -lean_inc(x_1); x_84 = l_Lean_getStructureCtor(x_83, x_1); x_85 = lean_ctor_get(x_84, 0); lean_inc(x_85); @@ -9696,9 +9692,8 @@ lean_dec(x_5); x_8 = lean_ctor_get(x_6, 0); lean_inc(x_8); lean_dec(x_6); -lean_inc(x_1); lean_inc(x_8); -x_9 = lean_environment_find(x_8, x_1); +x_9 = l_Lean_Environment_find_x3f(x_8, x_1); if (lean_obj_tag(x_9) == 0) { lean_object* x_10; diff --git a/stage0/stdlib/Lean/Elab/Quotation.c b/stage0/stdlib/Lean/Elab/Quotation.c index b09fdc00b280..17621a8be2af 100644 --- a/stage0/stdlib/Lean/Elab/Quotation.c +++ b/stage0/stdlib/Lean/Elab/Quotation.c @@ -918,7 +918,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30(lean_object*, lean_object*, 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_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__10; lean_object* l_Lean_Syntax_isLit_x3f(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot__1____x40_Lean_Elab_Quotation___hyg_6427____closed__1; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__7; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__39; @@ -1689,7 +1689,8 @@ x_12 = lean_ctor_get(x_10, 0); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); lean_dec(x_12); -x_14 = lean_environment_main_module(x_13); +x_14 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_15 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__3; x_16 = l_Lean_addMacroScope(x_14, x_15, x_9); x_17 = lean_box(0); @@ -1749,7 +1750,8 @@ lean_dec(x_10); x_36 = lean_ctor_get(x_34, 0); lean_inc(x_36); lean_dec(x_34); -x_37 = lean_environment_main_module(x_36); +x_37 = l_Lean_Environment_mainModule(x_36); +lean_dec(x_36); x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__3; x_39 = l_Lean_addMacroScope(x_37, x_38, x_9); x_40 = lean_box(0); @@ -2711,7 +2713,8 @@ x_24 = lean_ctor_get(x_22, 0); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); lean_dec(x_24); -x_26 = lean_environment_main_module(x_25); +x_26 = l_Lean_Environment_mainModule(x_25); +lean_dec(x_25); x_27 = l_Lean_Elab_Term_Quotation_mkTuple___closed__5; x_28 = l_Lean_addMacroScope(x_26, x_27, x_21); x_29 = l_Lean_Elab_Term_Quotation_mkTuple___closed__2; @@ -2763,7 +2766,8 @@ lean_dec(x_22); x_46 = lean_ctor_get(x_44, 0); lean_inc(x_46); lean_dec(x_44); -x_47 = lean_environment_main_module(x_46); +x_47 = l_Lean_Environment_mainModule(x_46); +lean_dec(x_46); x_48 = l_Lean_Elab_Term_Quotation_mkTuple___closed__5; x_49 = l_Lean_addMacroScope(x_47, x_48, x_21); x_50 = l_Lean_Elab_Term_Quotation_mkTuple___closed__2; @@ -2844,7 +2848,8 @@ x_75 = lean_ctor_get(x_73, 0); x_76 = lean_ctor_get(x_75, 0); lean_inc(x_76); lean_dec(x_75); -x_77 = lean_environment_main_module(x_76); +x_77 = l_Lean_Environment_mainModule(x_76); +lean_dec(x_76); x_78 = l_Lean_Elab_Term_Quotation_mkTuple___closed__14; x_79 = l_Lean_addMacroScope(x_77, x_78, x_72); x_80 = l_Lean_Elab_Term_Quotation_mkTuple___closed__11; @@ -2868,7 +2873,8 @@ lean_dec(x_73); x_85 = lean_ctor_get(x_83, 0); lean_inc(x_85); lean_dec(x_83); -x_86 = lean_environment_main_module(x_85); +x_86 = l_Lean_Environment_mainModule(x_85); +lean_dec(x_85); x_87 = l_Lean_Elab_Term_Quotation_mkTuple___closed__14; x_88 = l_Lean_addMacroScope(x_86, x_87, x_72); x_89 = l_Lean_Elab_Term_Quotation_mkTuple___closed__11; @@ -4219,7 +4225,8 @@ x_35 = lean_ctor_get(x_33, 0); x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); lean_dec(x_35); -x_37 = lean_environment_main_module(x_36); +x_37 = l_Lean_Environment_mainModule(x_36); +lean_dec(x_36); x_38 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__5; lean_inc(x_32); lean_inc(x_37); @@ -4262,7 +4269,8 @@ lean_dec(x_33); x_55 = lean_ctor_get(x_53, 0); lean_inc(x_55); lean_dec(x_53); -x_56 = lean_environment_main_module(x_55); +x_56 = l_Lean_Environment_mainModule(x_55); +lean_dec(x_55); x_57 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__5; lean_inc(x_32); lean_inc(x_56); @@ -4328,7 +4336,8 @@ x_88 = lean_ctor_get(x_86, 0); x_89 = lean_ctor_get(x_88, 0); lean_inc(x_89); lean_dec(x_88); -x_90 = lean_environment_main_module(x_89); +x_90 = l_Lean_Environment_mainModule(x_89); +lean_dec(x_89); x_91 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__17; lean_inc(x_85); lean_inc(x_90); @@ -4404,7 +4413,8 @@ lean_dec(x_86); x_118 = lean_ctor_get(x_116, 0); lean_inc(x_118); lean_dec(x_116); -x_119 = lean_environment_main_module(x_118); +x_119 = l_Lean_Environment_mainModule(x_118); +lean_dec(x_118); x_120 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__17; lean_inc(x_85); lean_inc(x_119); @@ -4504,7 +4514,8 @@ x_160 = lean_ctor_get(x_158, 0); x_161 = lean_ctor_get(x_160, 0); lean_inc(x_161); lean_dec(x_160); -x_162 = lean_environment_main_module(x_161); +x_162 = l_Lean_Environment_mainModule(x_161); +lean_dec(x_161); x_163 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__24; lean_inc(x_157); lean_inc(x_162); @@ -4577,7 +4588,8 @@ lean_dec(x_158); x_189 = lean_ctor_get(x_187, 0); lean_inc(x_189); lean_dec(x_187); -x_190 = lean_environment_main_module(x_189); +x_190 = l_Lean_Environment_mainModule(x_189); +lean_dec(x_189); x_191 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__24; lean_inc(x_157); lean_inc(x_190); @@ -4673,7 +4685,8 @@ x_229 = lean_ctor_get(x_227, 0); x_230 = lean_ctor_get(x_229, 0); lean_inc(x_230); lean_dec(x_229); -x_231 = lean_environment_main_module(x_230); +x_231 = l_Lean_Environment_mainModule(x_230); +lean_dec(x_230); x_232 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__31; lean_inc(x_226); lean_inc(x_231); @@ -4715,7 +4728,8 @@ lean_dec(x_227); x_248 = lean_ctor_get(x_246, 0); lean_inc(x_248); lean_dec(x_246); -x_249 = lean_environment_main_module(x_248); +x_249 = l_Lean_Environment_mainModule(x_248); +lean_dec(x_248); x_250 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__31; lean_inc(x_226); lean_inc(x_249); @@ -4779,7 +4793,8 @@ x_277 = lean_ctor_get(x_275, 0); x_278 = lean_ctor_get(x_277, 0); lean_inc(x_278); lean_dec(x_277); -x_279 = lean_environment_main_module(x_278); +x_279 = l_Lean_Environment_mainModule(x_278); +lean_dec(x_278); x_280 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__38; lean_inc(x_274); lean_inc(x_279); @@ -4821,7 +4836,8 @@ lean_dec(x_275); x_296 = lean_ctor_get(x_294, 0); lean_inc(x_296); lean_dec(x_294); -x_297 = lean_environment_main_module(x_296); +x_297 = l_Lean_Environment_mainModule(x_296); +lean_dec(x_296); x_298 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__38; lean_inc(x_274); lean_inc(x_297); @@ -4884,7 +4900,8 @@ x_324 = lean_ctor_get(x_322, 0); x_325 = lean_ctor_get(x_324, 0); lean_inc(x_325); lean_dec(x_324); -x_326 = lean_environment_main_module(x_325); +x_326 = l_Lean_Environment_mainModule(x_325); +lean_dec(x_325); x_327 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__45; lean_inc(x_321); lean_inc(x_326); @@ -4926,7 +4943,8 @@ lean_dec(x_322); x_343 = lean_ctor_get(x_341, 0); lean_inc(x_343); lean_dec(x_341); -x_344 = lean_environment_main_module(x_343); +x_344 = l_Lean_Environment_mainModule(x_343); +lean_dec(x_343); x_345 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__45; lean_inc(x_321); lean_inc(x_344); @@ -4988,7 +5006,8 @@ x_370 = lean_ctor_get(x_368, 0); x_371 = lean_ctor_get(x_370, 0); lean_inc(x_371); lean_dec(x_370); -x_372 = lean_environment_main_module(x_371); +x_372 = l_Lean_Environment_mainModule(x_371); +lean_dec(x_371); x_373 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__52; lean_inc(x_367); lean_inc(x_372); @@ -5030,7 +5049,8 @@ lean_dec(x_368); x_389 = lean_ctor_get(x_387, 0); lean_inc(x_389); lean_dec(x_387); -x_390 = lean_environment_main_module(x_389); +x_390 = l_Lean_Environment_mainModule(x_389); +lean_dec(x_389); x_391 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__52; lean_inc(x_367); lean_inc(x_390); @@ -5091,7 +5111,8 @@ x_415 = lean_ctor_get(x_413, 0); x_416 = lean_ctor_get(x_415, 0); lean_inc(x_416); lean_dec(x_415); -x_417 = lean_environment_main_module(x_416); +x_417 = l_Lean_Environment_mainModule(x_416); +lean_dec(x_416); x_418 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__59; lean_inc(x_412); lean_inc(x_417); @@ -5133,7 +5154,8 @@ lean_dec(x_413); x_434 = lean_ctor_get(x_432, 0); lean_inc(x_434); lean_dec(x_432); -x_435 = lean_environment_main_module(x_434); +x_435 = l_Lean_Environment_mainModule(x_434); +lean_dec(x_434); x_436 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__59; lean_inc(x_412); lean_inc(x_435); @@ -5193,7 +5215,8 @@ x_459 = lean_ctor_get(x_457, 0); x_460 = lean_ctor_get(x_459, 0); lean_inc(x_460); lean_dec(x_459); -x_461 = lean_environment_main_module(x_460); +x_461 = l_Lean_Environment_mainModule(x_460); +lean_dec(x_460); x_462 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__66; lean_inc(x_456); lean_inc(x_461); @@ -5235,7 +5258,8 @@ lean_dec(x_457); x_478 = lean_ctor_get(x_476, 0); lean_inc(x_478); lean_dec(x_476); -x_479 = lean_environment_main_module(x_478); +x_479 = l_Lean_Environment_mainModule(x_478); +lean_dec(x_478); x_480 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__66; lean_inc(x_456); lean_inc(x_479); @@ -5290,7 +5314,8 @@ x_503 = lean_ctor_get(x_501, 0); x_504 = lean_ctor_get(x_503, 0); lean_inc(x_504); lean_dec(x_503); -x_505 = lean_environment_main_module(x_504); +x_505 = l_Lean_Environment_mainModule(x_504); +lean_dec(x_504); x_506 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__5; lean_inc(x_500); lean_inc(x_505); @@ -5333,7 +5358,8 @@ lean_dec(x_501); x_523 = lean_ctor_get(x_521, 0); lean_inc(x_523); lean_dec(x_521); -x_524 = lean_environment_main_module(x_523); +x_524 = l_Lean_Environment_mainModule(x_523); +lean_dec(x_523); x_525 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__5; lean_inc(x_500); lean_inc(x_524); @@ -6330,7 +6356,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -6390,7 +6417,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -6450,7 +6478,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -6510,7 +6539,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -6570,7 +6600,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -6630,7 +6661,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -6918,7 +6950,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -6978,7 +7011,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -7038,7 +7072,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -7098,7 +7133,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -7158,7 +7194,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -7218,7 +7255,8 @@ lean_dec(x_21); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__4; x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); x_28 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; @@ -7789,7 +7827,8 @@ lean_dec(x_22); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_environment_main_module(x_25); +x_26 = l_Lean_Environment_mainModule(x_25); +lean_dec(x_25); x_27 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__2___closed__6; x_28 = l_Lean_addMacroScope(x_26, x_27, x_21); x_29 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__2___closed__5; @@ -8260,7 +8299,8 @@ x_43 = lean_ctor_get(x_40, 1); x_44 = lean_ctor_get(x_42, 0); lean_inc(x_44); lean_dec(x_42); -x_45 = lean_environment_main_module(x_44); +x_45 = l_Lean_Environment_mainModule(x_44); +lean_dec(x_44); x_46 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_47 = l_Lean_addMacroScope(x_45, x_46, x_39); x_48 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -8360,7 +8400,8 @@ lean_dec(x_40); x_86 = lean_ctor_get(x_84, 0); lean_inc(x_86); lean_dec(x_84); -x_87 = lean_environment_main_module(x_86); +x_87 = l_Lean_Environment_mainModule(x_86); +lean_dec(x_86); x_88 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_89 = l_Lean_addMacroScope(x_87, x_88, x_39); x_90 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -8481,7 +8522,8 @@ if (lean_is_exclusive(x_133)) { x_137 = lean_ctor_get(x_134, 0); lean_inc(x_137); lean_dec(x_134); -x_138 = lean_environment_main_module(x_137); +x_138 = l_Lean_Environment_mainModule(x_137); +lean_dec(x_137); x_139 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_140 = l_Lean_addMacroScope(x_138, x_139, x_132); x_141 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -8701,7 +8743,8 @@ x_218 = lean_ctor_get(x_215, 1); x_219 = lean_ctor_get(x_217, 0); lean_inc(x_219); lean_dec(x_217); -x_220 = lean_environment_main_module(x_219); +x_220 = l_Lean_Environment_mainModule(x_219); +lean_dec(x_219); x_221 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_222 = l_Lean_addMacroScope(x_220, x_221, x_214); x_223 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -8801,7 +8844,8 @@ lean_dec(x_215); x_261 = lean_ctor_get(x_259, 0); lean_inc(x_261); lean_dec(x_259); -x_262 = lean_environment_main_module(x_261); +x_262 = l_Lean_Environment_mainModule(x_261); +lean_dec(x_261); x_263 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_264 = l_Lean_addMacroScope(x_262, x_263, x_214); x_265 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -8922,7 +8966,8 @@ if (lean_is_exclusive(x_308)) { x_312 = lean_ctor_get(x_309, 0); lean_inc(x_312); lean_dec(x_309); -x_313 = lean_environment_main_module(x_312); +x_313 = l_Lean_Environment_mainModule(x_312); +lean_dec(x_312); x_314 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_315 = l_Lean_addMacroScope(x_313, x_314, x_307); x_316 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -9148,7 +9193,8 @@ x_396 = lean_ctor_get(x_393, 1); x_397 = lean_ctor_get(x_395, 0); lean_inc(x_397); lean_dec(x_395); -x_398 = lean_environment_main_module(x_397); +x_398 = l_Lean_Environment_mainModule(x_397); +lean_dec(x_397); x_399 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_400 = l_Lean_addMacroScope(x_398, x_399, x_392); x_401 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -9248,7 +9294,8 @@ lean_dec(x_393); x_439 = lean_ctor_get(x_437, 0); lean_inc(x_439); lean_dec(x_437); -x_440 = lean_environment_main_module(x_439); +x_440 = l_Lean_Environment_mainModule(x_439); +lean_dec(x_439); x_441 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_442 = l_Lean_addMacroScope(x_440, x_441, x_392); x_443 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -9369,7 +9416,8 @@ if (lean_is_exclusive(x_486)) { x_490 = lean_ctor_get(x_487, 0); lean_inc(x_490); lean_dec(x_487); -x_491 = lean_environment_main_module(x_490); +x_491 = l_Lean_Environment_mainModule(x_490); +lean_dec(x_490); x_492 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_493 = l_Lean_addMacroScope(x_491, x_492, x_485); x_494 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -9485,7 +9533,8 @@ x_552 = lean_ctor_get(x_549, 1); x_553 = lean_ctor_get(x_551, 0); lean_inc(x_553); lean_dec(x_551); -x_554 = lean_environment_main_module(x_553); +x_554 = l_Lean_Environment_mainModule(x_553); +lean_dec(x_553); x_555 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_547); lean_ctor_set_tag(x_549, 2); @@ -9604,7 +9653,8 @@ lean_dec(x_549); x_597 = lean_ctor_get(x_595, 0); lean_inc(x_597); lean_dec(x_595); -x_598 = lean_environment_main_module(x_597); +x_598 = l_Lean_Environment_mainModule(x_597); +lean_dec(x_597); x_599 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_547); x_600 = lean_alloc_ctor(2, 2, 0); @@ -9833,7 +9883,8 @@ x_665 = lean_ctor_get(x_662, 1); x_666 = lean_ctor_get(x_664, 0); lean_inc(x_666); lean_dec(x_664); -x_667 = lean_environment_main_module(x_666); +x_667 = l_Lean_Environment_mainModule(x_666); +lean_dec(x_666); x_668 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_669 = l_Lean_addMacroScope(x_667, x_668, x_661); x_670 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -9933,7 +9984,8 @@ lean_dec(x_662); x_708 = lean_ctor_get(x_706, 0); lean_inc(x_708); lean_dec(x_706); -x_709 = lean_environment_main_module(x_708); +x_709 = l_Lean_Environment_mainModule(x_708); +lean_dec(x_708); x_710 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_711 = l_Lean_addMacroScope(x_709, x_710, x_661); x_712 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -10054,7 +10106,8 @@ if (lean_is_exclusive(x_755)) { x_759 = lean_ctor_get(x_756, 0); lean_inc(x_759); lean_dec(x_756); -x_760 = lean_environment_main_module(x_759); +x_760 = l_Lean_Environment_mainModule(x_759); +lean_dec(x_759); x_761 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_762 = l_Lean_addMacroScope(x_760, x_761, x_754); x_763 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -10269,7 +10322,8 @@ x_839 = lean_ctor_get(x_836, 1); x_840 = lean_ctor_get(x_838, 0); lean_inc(x_840); lean_dec(x_838); -x_841 = lean_environment_main_module(x_840); +x_841 = l_Lean_Environment_mainModule(x_840); +lean_dec(x_840); x_842 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_843 = l_Lean_addMacroScope(x_841, x_842, x_835); x_844 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -10369,7 +10423,8 @@ lean_dec(x_836); x_882 = lean_ctor_get(x_880, 0); lean_inc(x_882); lean_dec(x_880); -x_883 = lean_environment_main_module(x_882); +x_883 = l_Lean_Environment_mainModule(x_882); +lean_dec(x_882); x_884 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_885 = l_Lean_addMacroScope(x_883, x_884, x_835); x_886 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -10490,7 +10545,8 @@ if (lean_is_exclusive(x_929)) { x_933 = lean_ctor_get(x_930, 0); lean_inc(x_933); lean_dec(x_930); -x_934 = lean_environment_main_module(x_933); +x_934 = l_Lean_Environment_mainModule(x_933); +lean_dec(x_933); x_935 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_936 = l_Lean_addMacroScope(x_934, x_935, x_928); x_937 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -10707,7 +10763,8 @@ x_1013 = lean_ctor_get(x_1010, 1); x_1014 = lean_ctor_get(x_1012, 0); lean_inc(x_1014); lean_dec(x_1012); -x_1015 = lean_environment_main_module(x_1014); +x_1015 = l_Lean_Environment_mainModule(x_1014); +lean_dec(x_1014); x_1016 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_1017 = l_Lean_addMacroScope(x_1015, x_1016, x_1009); x_1018 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -10807,7 +10864,8 @@ lean_dec(x_1010); x_1056 = lean_ctor_get(x_1054, 0); lean_inc(x_1056); lean_dec(x_1054); -x_1057 = lean_environment_main_module(x_1056); +x_1057 = l_Lean_Environment_mainModule(x_1056); +lean_dec(x_1056); x_1058 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_1059 = l_Lean_addMacroScope(x_1057, x_1058, x_1009); x_1060 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -10928,7 +10986,8 @@ if (lean_is_exclusive(x_1103)) { x_1107 = lean_ctor_get(x_1104, 0); lean_inc(x_1107); lean_dec(x_1104); -x_1108 = lean_environment_main_module(x_1107); +x_1108 = l_Lean_Environment_mainModule(x_1107); +lean_dec(x_1107); x_1109 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__4; x_1110 = l_Lean_addMacroScope(x_1108, x_1109, x_1102); x_1111 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__2; @@ -11869,7 +11928,8 @@ x_157 = lean_ctor_get(x_155, 0); x_158 = lean_ctor_get(x_157, 0); lean_inc(x_158); lean_dec(x_157); -x_159 = lean_environment_main_module(x_158); +x_159 = l_Lean_Environment_mainModule(x_158); +lean_dec(x_158); x_160 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_153); x_161 = lean_alloc_ctor(2, 2, 0); @@ -11908,7 +11968,8 @@ lean_dec(x_155); x_175 = lean_ctor_get(x_173, 0); lean_inc(x_175); lean_dec(x_173); -x_176 = lean_environment_main_module(x_175); +x_176 = l_Lean_Environment_mainModule(x_175); +lean_dec(x_175); x_177 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_153); x_178 = lean_alloc_ctor(2, 2, 0); @@ -11962,7 +12023,8 @@ x_197 = lean_ctor_get(x_195, 0); x_198 = lean_ctor_get(x_197, 0); lean_inc(x_198); lean_dec(x_197); -x_199 = lean_environment_main_module(x_198); +x_199 = l_Lean_Environment_mainModule(x_198); +lean_dec(x_198); x_200 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_193); x_201 = lean_alloc_ctor(2, 2, 0); @@ -12001,7 +12063,8 @@ lean_dec(x_195); x_215 = lean_ctor_get(x_213, 0); lean_inc(x_215); lean_dec(x_213); -x_216 = lean_environment_main_module(x_215); +x_216 = l_Lean_Environment_mainModule(x_215); +lean_dec(x_215); x_217 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_193); x_218 = lean_alloc_ctor(2, 2, 0); @@ -12055,7 +12118,8 @@ x_237 = lean_ctor_get(x_235, 0); x_238 = lean_ctor_get(x_237, 0); lean_inc(x_238); lean_dec(x_237); -x_239 = lean_environment_main_module(x_238); +x_239 = l_Lean_Environment_mainModule(x_238); +lean_dec(x_238); x_240 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_233); x_241 = lean_alloc_ctor(2, 2, 0); @@ -12245,7 +12309,8 @@ lean_dec(x_235); x_320 = lean_ctor_get(x_318, 0); lean_inc(x_320); lean_dec(x_318); -x_321 = lean_environment_main_module(x_320); +x_321 = l_Lean_Environment_mainModule(x_320); +lean_dec(x_320); x_322 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_233); x_323 = lean_alloc_ctor(2, 2, 0); @@ -14425,7 +14490,8 @@ x_222 = lean_ctor_get(x_220, 0); x_223 = lean_ctor_get(x_222, 0); lean_inc(x_223); lean_dec(x_222); -x_224 = lean_environment_main_module(x_223); +x_224 = l_Lean_Environment_mainModule(x_223); +lean_dec(x_223); x_225 = lean_box(0); x_226 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_218); @@ -14468,7 +14534,8 @@ lean_dec(x_220); x_244 = lean_ctor_get(x_242, 0); lean_inc(x_244); lean_dec(x_242); -x_245 = lean_environment_main_module(x_244); +x_245 = l_Lean_Environment_mainModule(x_244); +lean_dec(x_244); x_246 = lean_box(0); x_247 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_218); @@ -14770,7 +14837,8 @@ x_69 = lean_ctor_get(x_67, 0); x_70 = lean_ctor_get(x_69, 0); lean_inc(x_70); lean_dec(x_69); -x_71 = lean_environment_main_module(x_70); +x_71 = l_Lean_Environment_mainModule(x_70); +lean_dec(x_70); x_72 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; lean_inc(x_66); lean_inc(x_71); @@ -14867,7 +14935,8 @@ lean_dec(x_67); x_112 = lean_ctor_get(x_110, 0); lean_inc(x_112); lean_dec(x_110); -x_113 = lean_environment_main_module(x_112); +x_113 = l_Lean_Environment_mainModule(x_112); +lean_dec(x_112); x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; lean_inc(x_66); lean_inc(x_113); @@ -14987,7 +15056,8 @@ if (lean_is_exclusive(x_158)) { x_162 = lean_ctor_get(x_159, 0); lean_inc(x_162); lean_dec(x_159); -x_163 = lean_environment_main_module(x_162); +x_163 = l_Lean_Environment_mainModule(x_162); +lean_dec(x_162); x_164 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; lean_inc(x_157); lean_inc(x_163); @@ -15427,7 +15497,8 @@ x_36 = lean_ctor_get(x_33, 0); x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); lean_dec(x_36); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__4; lean_inc(x_32); lean_inc(x_38); @@ -15538,7 +15609,8 @@ lean_dec(x_33); x_86 = lean_ctor_get(x_84, 0); lean_inc(x_86); lean_dec(x_84); -x_87 = lean_environment_main_module(x_86); +x_87 = l_Lean_Environment_mainModule(x_86); +lean_dec(x_86); x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__4; lean_inc(x_32); lean_inc(x_87); @@ -15837,7 +15909,8 @@ x_35 = lean_ctor_get(x_33, 0); x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); lean_dec(x_35); -x_37 = lean_environment_main_module(x_36); +x_37 = l_Lean_Environment_mainModule(x_36); +lean_dec(x_36); x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; lean_inc(x_32); lean_inc(x_37); @@ -15882,7 +15955,8 @@ lean_dec(x_33); x_56 = lean_ctor_get(x_54, 0); lean_inc(x_56); lean_dec(x_54); -x_57 = lean_environment_main_module(x_56); +x_57 = l_Lean_Environment_mainModule(x_56); +lean_dec(x_56); x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; lean_inc(x_32); lean_inc(x_57); @@ -15961,7 +16035,8 @@ x_89 = lean_ctor_get(x_87, 0); x_90 = lean_ctor_get(x_89, 0); lean_inc(x_90); lean_dec(x_89); -x_91 = lean_environment_main_module(x_90); +x_91 = l_Lean_Environment_mainModule(x_90); +lean_dec(x_90); x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__4; lean_inc(x_86); lean_inc(x_91); @@ -16128,7 +16203,8 @@ lean_dec(x_87); x_163 = lean_ctor_get(x_161, 0); lean_inc(x_163); lean_dec(x_161); -x_164 = lean_environment_main_module(x_163); +x_164 = l_Lean_Environment_mainModule(x_163); +lean_dec(x_163); x_165 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__4; lean_inc(x_86); lean_inc(x_164); @@ -16296,7 +16372,8 @@ if (lean_is_exclusive(x_220)) { x_224 = lean_ctor_get(x_221, 0); lean_inc(x_224); lean_dec(x_221); -x_225 = lean_environment_main_module(x_224); +x_225 = l_Lean_Environment_mainModule(x_224); +lean_dec(x_224); x_226 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__4; lean_inc(x_219); lean_inc(x_225); @@ -18382,7 +18459,8 @@ if (lean_is_exclusive(x_14)) { x_18 = lean_ctor_get(x_15, 0); lean_inc(x_18); lean_dec(x_15); -x_19 = lean_environment_main_module(x_18); +x_19 = l_Lean_Environment_mainModule(x_18); +lean_dec(x_18); x_20 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__5; lean_inc(x_13); lean_inc(x_19); @@ -20905,7 +20983,8 @@ lean_dec(x_20); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; lean_inc(x_19); lean_inc(x_24); @@ -20972,7 +21051,8 @@ lean_dec(x_49); x_52 = lean_ctor_get(x_50, 0); lean_inc(x_52); lean_dec(x_50); -x_53 = lean_environment_main_module(x_52); +x_53 = l_Lean_Environment_mainModule(x_52); +lean_dec(x_52); x_54 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; lean_inc(x_48); lean_inc(x_53); @@ -21168,7 +21248,8 @@ lean_dec(x_20); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; lean_inc(x_19); lean_inc(x_24); @@ -21235,7 +21316,8 @@ lean_dec(x_49); x_52 = lean_ctor_get(x_50, 0); lean_inc(x_52); lean_dec(x_50); -x_53 = lean_environment_main_module(x_52); +x_53 = l_Lean_Environment_mainModule(x_52); +lean_dec(x_52); x_54 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; lean_inc(x_48); lean_inc(x_53); @@ -21326,7 +21408,8 @@ x_35 = lean_ctor_get(x_32, 1); x_36 = lean_ctor_get(x_34, 0); lean_inc(x_36); lean_dec(x_34); -x_37 = lean_environment_main_module(x_36); +x_37 = l_Lean_Environment_mainModule(x_36); +lean_dec(x_36); x_38 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; lean_inc(x_31); lean_inc(x_37); @@ -21428,7 +21511,8 @@ lean_dec(x_32); x_68 = lean_ctor_get(x_66, 0); lean_inc(x_68); lean_dec(x_66); -x_69 = lean_environment_main_module(x_68); +x_69 = l_Lean_Environment_mainModule(x_68); +lean_dec(x_68); x_70 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; lean_inc(x_31); lean_inc(x_69); @@ -21551,7 +21635,8 @@ if (lean_is_exclusive(x_105)) { x_109 = lean_ctor_get(x_106, 0); lean_inc(x_109); lean_dec(x_106); -x_110 = lean_environment_main_module(x_109); +x_110 = l_Lean_Environment_mainModule(x_109); +lean_dec(x_109); x_111 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; lean_inc(x_104); lean_inc(x_110); @@ -21805,7 +21890,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); lean_ctor_set_tag(x_26, 2); @@ -21909,7 +21995,8 @@ lean_dec(x_26); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); x_76 = lean_alloc_ctor(2, 2, 0); @@ -22044,7 +22131,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); lean_ctor_set_tag(x_26, 2); @@ -22148,7 +22236,8 @@ lean_dec(x_26); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); x_76 = lean_alloc_ctor(2, 2, 0); @@ -22283,7 +22372,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); lean_ctor_set_tag(x_26, 2); @@ -22387,7 +22477,8 @@ lean_dec(x_26); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); x_76 = lean_alloc_ctor(2, 2, 0); @@ -22610,7 +22701,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); lean_ctor_set_tag(x_26, 2); @@ -22714,7 +22806,8 @@ lean_dec(x_26); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); x_76 = lean_alloc_ctor(2, 2, 0); @@ -22849,7 +22942,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); lean_ctor_set_tag(x_26, 2); @@ -22953,7 +23047,8 @@ lean_dec(x_26); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); x_76 = lean_alloc_ctor(2, 2, 0); @@ -23088,7 +23183,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); lean_ctor_set_tag(x_26, 2); @@ -23192,7 +23288,8 @@ lean_dec(x_26); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_24); x_76 = lean_alloc_ctor(2, 2, 0); @@ -23411,7 +23508,8 @@ x_24 = lean_ctor_get(x_21, 1); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_environment_main_module(x_25); +x_26 = l_Lean_Environment_mainModule(x_25); +lean_dec(x_25); x_27 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__3; lean_inc(x_20); lean_inc(x_26); @@ -23548,7 +23646,8 @@ lean_dec(x_21); x_81 = lean_ctor_get(x_79, 0); lean_inc(x_81); lean_dec(x_79); -x_82 = lean_environment_main_module(x_81); +x_82 = l_Lean_Environment_mainModule(x_81); +lean_dec(x_81); x_83 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__3; lean_inc(x_20); lean_inc(x_82); @@ -23706,7 +23805,8 @@ if (lean_is_exclusive(x_142)) { x_146 = lean_ctor_get(x_143, 0); lean_inc(x_146); lean_dec(x_143); -x_147 = lean_environment_main_module(x_146); +x_147 = l_Lean_Environment_mainModule(x_146); +lean_dec(x_146); x_148 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__3; lean_inc(x_141); lean_inc(x_147); @@ -23862,7 +23962,8 @@ x_17 = lean_ctor_get(x_15, 0); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); lean_dec(x_17); -x_19 = lean_environment_main_module(x_18); +x_19 = l_Lean_Environment_mainModule(x_18); +lean_dec(x_18); x_20 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_13); x_21 = lean_alloc_ctor(2, 2, 0); @@ -23920,7 +24021,8 @@ lean_dec(x_15); x_43 = lean_ctor_get(x_41, 0); lean_inc(x_43); lean_dec(x_41); -x_44 = lean_environment_main_module(x_43); +x_44 = l_Lean_Environment_mainModule(x_43); +lean_dec(x_43); x_45 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_13); x_46 = lean_alloc_ctor(2, 2, 0); @@ -24295,7 +24397,8 @@ x_27 = lean_ctor_get(x_25, 0); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); lean_dec(x_27); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; lean_inc(x_24); lean_inc(x_29); @@ -24391,7 +24494,8 @@ lean_dec(x_25); x_64 = lean_ctor_get(x_62, 0); lean_inc(x_64); lean_dec(x_62); -x_65 = lean_environment_main_module(x_64); +x_65 = l_Lean_Environment_mainModule(x_64); +lean_dec(x_64); x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; lean_inc(x_24); lean_inc(x_65); @@ -24600,7 +24704,8 @@ if (lean_is_exclusive(x_119)) { x_123 = lean_ctor_get(x_120, 0); lean_inc(x_123); lean_dec(x_120); -x_124 = lean_environment_main_module(x_123); +x_124 = l_Lean_Environment_mainModule(x_123); +lean_dec(x_123); x_125 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; lean_inc(x_118); lean_inc(x_124); @@ -24989,7 +25094,8 @@ x_30 = lean_ctor_get(x_27, 1); x_31 = lean_ctor_get(x_29, 0); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_environment_main_module(x_31); +x_32 = l_Lean_Environment_mainModule(x_31); +lean_dec(x_31); x_33 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__8; lean_inc(x_26); lean_inc(x_32); @@ -25079,7 +25185,8 @@ lean_dec(x_27); x_68 = lean_ctor_get(x_66, 0); lean_inc(x_68); lean_dec(x_66); -x_69 = lean_environment_main_module(x_68); +x_69 = l_Lean_Environment_mainModule(x_68); +lean_dec(x_68); x_70 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__8; lean_inc(x_26); lean_inc(x_69); @@ -25180,7 +25287,8 @@ x_111 = lean_ctor_get(x_108, 1); x_112 = lean_ctor_get(x_110, 0); lean_inc(x_112); lean_dec(x_110); -x_113 = lean_environment_main_module(x_112); +x_113 = l_Lean_Environment_mainModule(x_112); +lean_dec(x_112); x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4; lean_inc(x_107); lean_inc(x_113); @@ -25271,7 +25379,8 @@ lean_dec(x_108); x_150 = lean_ctor_get(x_148, 0); lean_inc(x_150); lean_dec(x_148); -x_151 = lean_environment_main_module(x_150); +x_151 = l_Lean_Environment_mainModule(x_150); +lean_dec(x_150); x_152 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4; lean_inc(x_107); lean_inc(x_151); @@ -25375,7 +25484,8 @@ x_194 = lean_ctor_get(x_191, 1); x_195 = lean_ctor_get(x_193, 0); lean_inc(x_195); lean_dec(x_193); -x_196 = lean_environment_main_module(x_195); +x_196 = l_Lean_Environment_mainModule(x_195); +lean_dec(x_195); x_197 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9; lean_inc(x_190); lean_inc(x_196); @@ -25429,7 +25539,8 @@ lean_dec(x_191); x_218 = lean_ctor_get(x_216, 0); lean_inc(x_218); lean_dec(x_216); -x_219 = lean_environment_main_module(x_218); +x_219 = l_Lean_Environment_mainModule(x_218); +lean_dec(x_218); x_220 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9; lean_inc(x_190); lean_inc(x_219); @@ -25495,7 +25606,8 @@ x_247 = lean_ctor_get(x_244, 1); x_248 = lean_ctor_get(x_246, 0); lean_inc(x_248); lean_dec(x_246); -x_249 = lean_environment_main_module(x_248); +x_249 = l_Lean_Environment_mainModule(x_248); +lean_dec(x_248); x_250 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__8; lean_inc(x_243); lean_inc(x_249); @@ -25585,7 +25697,8 @@ lean_dec(x_244); x_285 = lean_ctor_get(x_283, 0); lean_inc(x_285); lean_dec(x_283); -x_286 = lean_environment_main_module(x_285); +x_286 = l_Lean_Environment_mainModule(x_285); +lean_dec(x_285); x_287 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__8; lean_inc(x_243); lean_inc(x_286); @@ -25686,7 +25799,8 @@ x_328 = lean_ctor_get(x_325, 1); x_329 = lean_ctor_get(x_327, 0); lean_inc(x_329); lean_dec(x_327); -x_330 = lean_environment_main_module(x_329); +x_330 = l_Lean_Environment_mainModule(x_329); +lean_dec(x_329); x_331 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__8; lean_inc(x_324); lean_inc(x_330); @@ -25776,7 +25890,8 @@ lean_dec(x_325); x_366 = lean_ctor_get(x_364, 0); lean_inc(x_366); lean_dec(x_364); -x_367 = lean_environment_main_module(x_366); +x_367 = l_Lean_Environment_mainModule(x_366); +lean_dec(x_366); x_368 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__8; lean_inc(x_324); lean_inc(x_367); @@ -25877,7 +25992,8 @@ x_409 = lean_ctor_get(x_406, 1); x_410 = lean_ctor_get(x_408, 0); lean_inc(x_410); lean_dec(x_408); -x_411 = lean_environment_main_module(x_410); +x_411 = l_Lean_Environment_mainModule(x_410); +lean_dec(x_410); x_412 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__14; lean_inc(x_405); lean_inc(x_411); @@ -26038,7 +26154,8 @@ lean_dec(x_406); x_476 = lean_ctor_get(x_474, 0); lean_inc(x_476); lean_dec(x_474); -x_477 = lean_environment_main_module(x_476); +x_477 = l_Lean_Environment_mainModule(x_476); +lean_dec(x_476); x_478 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__14; lean_inc(x_405); lean_inc(x_477); @@ -26671,7 +26788,8 @@ x_27 = lean_ctor_get(x_24, 1); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__17; lean_inc(x_23); lean_inc(x_29); @@ -26830,7 +26948,8 @@ x_86 = lean_ctor_get(x_84, 0); x_87 = lean_ctor_get(x_86, 0); lean_inc(x_87); lean_dec(x_86); -x_88 = lean_environment_main_module(x_87); +x_88 = l_Lean_Environment_mainModule(x_87); +lean_dec(x_87); x_89 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; lean_inc(x_23); lean_inc(x_88); @@ -26901,7 +27020,8 @@ lean_dec(x_84); x_112 = lean_ctor_get(x_110, 0); lean_inc(x_112); lean_dec(x_110); -x_113 = lean_environment_main_module(x_112); +x_113 = l_Lean_Environment_mainModule(x_112); +lean_dec(x_112); x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; lean_inc(x_23); lean_inc(x_113); @@ -27085,7 +27205,8 @@ if (lean_is_exclusive(x_154)) { x_158 = lean_ctor_get(x_155, 0); lean_inc(x_158); lean_dec(x_155); -x_159 = lean_environment_main_module(x_158); +x_159 = l_Lean_Environment_mainModule(x_158); +lean_dec(x_158); x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; lean_inc(x_23); lean_inc(x_159); @@ -27234,7 +27355,8 @@ lean_dec(x_24); x_192 = lean_ctor_get(x_190, 0); lean_inc(x_192); lean_dec(x_190); -x_193 = lean_environment_main_module(x_192); +x_193 = l_Lean_Environment_mainModule(x_192); +lean_dec(x_192); x_194 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__17; lean_inc(x_23); lean_inc(x_193); @@ -27411,7 +27533,8 @@ if (lean_is_exclusive(x_250)) { x_254 = lean_ctor_get(x_251, 0); lean_inc(x_254); lean_dec(x_251); -x_255 = lean_environment_main_module(x_254); +x_255 = l_Lean_Environment_mainModule(x_254); +lean_dec(x_254); x_256 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; lean_inc(x_23); lean_inc(x_255); @@ -27580,7 +27703,8 @@ if (lean_is_exclusive(x_292)) { x_296 = lean_ctor_get(x_293, 0); lean_inc(x_296); lean_dec(x_293); -x_297 = lean_environment_main_module(x_296); +x_297 = l_Lean_Environment_mainModule(x_296); +lean_dec(x_296); x_298 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__17; lean_inc(x_291); lean_inc(x_297); @@ -27762,7 +27886,8 @@ if (lean_is_exclusive(x_355)) { x_359 = lean_ctor_get(x_356, 0); lean_inc(x_359); lean_dec(x_356); -x_360 = lean_environment_main_module(x_359); +x_360 = l_Lean_Environment_mainModule(x_359); +lean_dec(x_359); x_361 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; lean_inc(x_291); lean_inc(x_360); @@ -28073,7 +28198,8 @@ x_27 = lean_ctor_get(x_24, 0); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); lean_dec(x_27); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_22); x_31 = lean_alloc_ctor(2, 2, 0); @@ -28299,7 +28425,8 @@ lean_dec(x_24); x_116 = lean_ctor_get(x_114, 0); lean_inc(x_116); lean_dec(x_114); -x_117 = lean_environment_main_module(x_116); +x_117 = l_Lean_Environment_mainModule(x_116); +lean_dec(x_116); x_118 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_22); x_119 = lean_alloc_ctor(2, 2, 0); @@ -28588,7 +28715,8 @@ lean_dec(x_34); x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; x_40 = l_Lean_addMacroScope(x_38, x_39, x_33); x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__1; @@ -28662,7 +28790,8 @@ lean_dec(x_34); x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; x_40 = l_Lean_addMacroScope(x_38, x_39, x_33); x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__1; @@ -28736,7 +28865,8 @@ lean_dec(x_34); x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; x_40 = l_Lean_addMacroScope(x_38, x_39, x_33); x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__1; @@ -28810,7 +28940,8 @@ lean_dec(x_34); x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; x_40 = l_Lean_addMacroScope(x_38, x_39, x_33); x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__1; @@ -28884,7 +29015,8 @@ lean_dec(x_34); x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; x_40 = l_Lean_addMacroScope(x_38, x_39, x_33); x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__1; @@ -28958,7 +29090,8 @@ lean_dec(x_34); x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; x_40 = l_Lean_addMacroScope(x_38, x_39, x_33); x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__1; @@ -29270,7 +29403,8 @@ x_31 = lean_ctor_get(x_28, 1); x_32 = lean_ctor_get(x_30, 0); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_environment_main_module(x_32); +x_33 = l_Lean_Environment_mainModule(x_32); +lean_dec(x_32); x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_27); lean_inc(x_33); @@ -29337,7 +29471,8 @@ x_55 = lean_ctor_get(x_52, 1); x_56 = lean_ctor_get(x_54, 0); lean_inc(x_56); lean_dec(x_54); -x_57 = lean_environment_main_module(x_56); +x_57 = l_Lean_Environment_mainModule(x_56); +lean_dec(x_56); x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_59 = l_Lean_addMacroScope(x_57, x_58, x_27); lean_inc(x_2); @@ -29378,7 +29513,8 @@ lean_dec(x_52); x_69 = lean_ctor_get(x_67, 0); lean_inc(x_69); lean_dec(x_67); -x_70 = lean_environment_main_module(x_69); +x_70 = l_Lean_Environment_mainModule(x_69); +lean_dec(x_69); x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_72 = l_Lean_addMacroScope(x_70, x_71, x_27); lean_inc(x_2); @@ -29421,7 +29557,8 @@ lean_dec(x_28); x_83 = lean_ctor_get(x_81, 0); lean_inc(x_83); lean_dec(x_81); -x_84 = lean_environment_main_module(x_83); +x_84 = l_Lean_Environment_mainModule(x_83); +lean_dec(x_83); x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_27); lean_inc(x_84); @@ -29494,7 +29631,8 @@ if (lean_is_exclusive(x_104)) { x_108 = lean_ctor_get(x_105, 0); lean_inc(x_108); lean_dec(x_105); -x_109 = lean_environment_main_module(x_108); +x_109 = l_Lean_Environment_mainModule(x_108); +lean_dec(x_108); x_110 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_111 = l_Lean_addMacroScope(x_109, x_110, x_27); lean_inc(x_2); @@ -29562,7 +29700,8 @@ x_130 = lean_ctor_get(x_127, 1); x_131 = lean_ctor_get(x_129, 0); lean_inc(x_131); lean_dec(x_129); -x_132 = lean_environment_main_module(x_131); +x_132 = l_Lean_Environment_mainModule(x_131); +lean_dec(x_131); x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_126); lean_inc(x_132); @@ -29630,7 +29769,8 @@ x_154 = lean_ctor_get(x_151, 1); x_155 = lean_ctor_get(x_153, 0); lean_inc(x_155); lean_dec(x_153); -x_156 = lean_environment_main_module(x_155); +x_156 = l_Lean_Environment_mainModule(x_155); +lean_dec(x_155); x_157 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_158 = l_Lean_addMacroScope(x_156, x_157, x_126); lean_inc(x_2); @@ -29671,7 +29811,8 @@ lean_dec(x_151); x_168 = lean_ctor_get(x_166, 0); lean_inc(x_168); lean_dec(x_166); -x_169 = lean_environment_main_module(x_168); +x_169 = l_Lean_Environment_mainModule(x_168); +lean_dec(x_168); x_170 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_171 = l_Lean_addMacroScope(x_169, x_170, x_126); lean_inc(x_2); @@ -29714,7 +29855,8 @@ lean_dec(x_127); x_182 = lean_ctor_get(x_180, 0); lean_inc(x_182); lean_dec(x_180); -x_183 = lean_environment_main_module(x_182); +x_183 = l_Lean_Environment_mainModule(x_182); +lean_dec(x_182); x_184 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_126); lean_inc(x_183); @@ -29788,7 +29930,8 @@ if (lean_is_exclusive(x_203)) { x_207 = lean_ctor_get(x_204, 0); lean_inc(x_207); lean_dec(x_204); -x_208 = lean_environment_main_module(x_207); +x_208 = l_Lean_Environment_mainModule(x_207); +lean_dec(x_207); x_209 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_210 = l_Lean_addMacroScope(x_208, x_209, x_126); lean_inc(x_2); @@ -29865,7 +30008,8 @@ x_232 = lean_ctor_get(x_229, 1); x_233 = lean_ctor_get(x_231, 0); lean_inc(x_233); lean_dec(x_231); -x_234 = lean_environment_main_module(x_233); +x_234 = l_Lean_Environment_mainModule(x_233); +lean_dec(x_233); x_235 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_228); lean_inc(x_234); @@ -29933,7 +30077,8 @@ x_256 = lean_ctor_get(x_253, 1); x_257 = lean_ctor_get(x_255, 0); lean_inc(x_257); lean_dec(x_255); -x_258 = lean_environment_main_module(x_257); +x_258 = l_Lean_Environment_mainModule(x_257); +lean_dec(x_257); x_259 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_260 = l_Lean_addMacroScope(x_258, x_259, x_228); lean_inc(x_2); @@ -29974,7 +30119,8 @@ lean_dec(x_253); x_270 = lean_ctor_get(x_268, 0); lean_inc(x_270); lean_dec(x_268); -x_271 = lean_environment_main_module(x_270); +x_271 = l_Lean_Environment_mainModule(x_270); +lean_dec(x_270); x_272 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_273 = l_Lean_addMacroScope(x_271, x_272, x_228); lean_inc(x_2); @@ -30017,7 +30163,8 @@ lean_dec(x_229); x_284 = lean_ctor_get(x_282, 0); lean_inc(x_284); lean_dec(x_282); -x_285 = lean_environment_main_module(x_284); +x_285 = l_Lean_Environment_mainModule(x_284); +lean_dec(x_284); x_286 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_228); lean_inc(x_285); @@ -30091,7 +30238,8 @@ if (lean_is_exclusive(x_305)) { x_309 = lean_ctor_get(x_306, 0); lean_inc(x_309); lean_dec(x_306); -x_310 = lean_environment_main_module(x_309); +x_310 = l_Lean_Environment_mainModule(x_309); +lean_dec(x_309); x_311 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_312 = l_Lean_addMacroScope(x_310, x_311, x_228); lean_inc(x_2); @@ -30154,7 +30302,8 @@ x_328 = lean_ctor_get(x_325, 1); x_329 = lean_ctor_get(x_327, 0); lean_inc(x_329); lean_dec(x_327); -x_330 = lean_environment_main_module(x_329); +x_330 = l_Lean_Environment_mainModule(x_329); +lean_dec(x_329); x_331 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__64; lean_inc(x_324); x_332 = l_Lean_addMacroScope(x_330, x_331, x_324); @@ -30186,7 +30335,8 @@ x_341 = lean_ctor_get(x_339, 0); x_342 = lean_ctor_get(x_341, 0); lean_inc(x_342); lean_dec(x_341); -x_343 = lean_environment_main_module(x_342); +x_343 = l_Lean_Environment_mainModule(x_342); +lean_dec(x_342); x_344 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9; lean_inc(x_323); x_345 = lean_alloc_ctor(2, 2, 0); @@ -30414,7 +30564,8 @@ lean_dec(x_339); x_438 = lean_ctor_get(x_436, 0); lean_inc(x_438); lean_dec(x_436); -x_439 = lean_environment_main_module(x_438); +x_439 = l_Lean_Environment_mainModule(x_438); +lean_dec(x_438); x_440 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9; lean_inc(x_323); x_441 = lean_alloc_ctor(2, 2, 0); @@ -30645,7 +30796,8 @@ lean_dec(x_325); x_535 = lean_ctor_get(x_533, 0); lean_inc(x_535); lean_dec(x_533); -x_536 = lean_environment_main_module(x_535); +x_536 = l_Lean_Environment_mainModule(x_535); +lean_dec(x_535); x_537 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__64; lean_inc(x_324); x_538 = l_Lean_addMacroScope(x_536, x_537, x_324); @@ -30684,7 +30836,8 @@ if (lean_is_exclusive(x_546)) { x_550 = lean_ctor_get(x_547, 0); lean_inc(x_550); lean_dec(x_547); -x_551 = lean_environment_main_module(x_550); +x_551 = l_Lean_Environment_mainModule(x_550); +lean_dec(x_550); x_552 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9; lean_inc(x_323); x_553 = lean_alloc_ctor(2, 2, 0); @@ -30936,7 +31089,8 @@ x_654 = lean_ctor_get(x_651, 1); x_655 = lean_ctor_get(x_653, 0); lean_inc(x_655); lean_dec(x_653); -x_656 = lean_environment_main_module(x_655); +x_656 = l_Lean_Environment_mainModule(x_655); +lean_dec(x_655); x_657 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_650); lean_inc(x_656); @@ -31004,7 +31158,8 @@ x_678 = lean_ctor_get(x_675, 1); x_679 = lean_ctor_get(x_677, 0); lean_inc(x_679); lean_dec(x_677); -x_680 = lean_environment_main_module(x_679); +x_680 = l_Lean_Environment_mainModule(x_679); +lean_dec(x_679); x_681 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_682 = l_Lean_addMacroScope(x_680, x_681, x_650); lean_inc(x_2); @@ -31045,7 +31200,8 @@ lean_dec(x_675); x_692 = lean_ctor_get(x_690, 0); lean_inc(x_692); lean_dec(x_690); -x_693 = lean_environment_main_module(x_692); +x_693 = l_Lean_Environment_mainModule(x_692); +lean_dec(x_692); x_694 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_695 = l_Lean_addMacroScope(x_693, x_694, x_650); lean_inc(x_2); @@ -31088,7 +31244,8 @@ lean_dec(x_651); x_706 = lean_ctor_get(x_704, 0); lean_inc(x_706); lean_dec(x_704); -x_707 = lean_environment_main_module(x_706); +x_707 = l_Lean_Environment_mainModule(x_706); +lean_dec(x_706); x_708 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_650); lean_inc(x_707); @@ -31162,7 +31319,8 @@ if (lean_is_exclusive(x_727)) { x_731 = lean_ctor_get(x_728, 0); lean_inc(x_731); lean_dec(x_728); -x_732 = lean_environment_main_module(x_731); +x_732 = l_Lean_Environment_mainModule(x_731); +lean_dec(x_731); x_733 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_734 = l_Lean_addMacroScope(x_732, x_733, x_650); lean_inc(x_2); @@ -31226,7 +31384,8 @@ x_752 = lean_ctor_get(x_749, 1); x_753 = lean_ctor_get(x_751, 0); lean_inc(x_753); lean_dec(x_751); -x_754 = lean_environment_main_module(x_753); +x_754 = l_Lean_Environment_mainModule(x_753); +lean_dec(x_753); x_755 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_748); lean_inc(x_754); @@ -31294,7 +31453,8 @@ x_776 = lean_ctor_get(x_773, 1); x_777 = lean_ctor_get(x_775, 0); lean_inc(x_777); lean_dec(x_775); -x_778 = lean_environment_main_module(x_777); +x_778 = l_Lean_Environment_mainModule(x_777); +lean_dec(x_777); x_779 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_780 = l_Lean_addMacroScope(x_778, x_779, x_748); lean_inc(x_2); @@ -31335,7 +31495,8 @@ lean_dec(x_773); x_790 = lean_ctor_get(x_788, 0); lean_inc(x_790); lean_dec(x_788); -x_791 = lean_environment_main_module(x_790); +x_791 = l_Lean_Environment_mainModule(x_790); +lean_dec(x_790); x_792 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_793 = l_Lean_addMacroScope(x_791, x_792, x_748); lean_inc(x_2); @@ -31378,7 +31539,8 @@ lean_dec(x_749); x_804 = lean_ctor_get(x_802, 0); lean_inc(x_804); lean_dec(x_802); -x_805 = lean_environment_main_module(x_804); +x_805 = l_Lean_Environment_mainModule(x_804); +lean_dec(x_804); x_806 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_748); lean_inc(x_805); @@ -31452,7 +31614,8 @@ if (lean_is_exclusive(x_825)) { x_829 = lean_ctor_get(x_826, 0); lean_inc(x_829); lean_dec(x_826); -x_830 = lean_environment_main_module(x_829); +x_830 = l_Lean_Environment_mainModule(x_829); +lean_dec(x_829); x_831 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_832 = l_Lean_addMacroScope(x_830, x_831, x_748); lean_inc(x_2); @@ -31517,7 +31680,8 @@ x_850 = lean_ctor_get(x_847, 1); x_851 = lean_ctor_get(x_849, 0); lean_inc(x_851); lean_dec(x_849); -x_852 = lean_environment_main_module(x_851); +x_852 = l_Lean_Environment_mainModule(x_851); +lean_dec(x_851); x_853 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_846); lean_inc(x_852); @@ -31585,7 +31749,8 @@ x_874 = lean_ctor_get(x_871, 1); x_875 = lean_ctor_get(x_873, 0); lean_inc(x_875); lean_dec(x_873); -x_876 = lean_environment_main_module(x_875); +x_876 = l_Lean_Environment_mainModule(x_875); +lean_dec(x_875); x_877 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_878 = l_Lean_addMacroScope(x_876, x_877, x_846); lean_inc(x_2); @@ -31626,7 +31791,8 @@ lean_dec(x_871); x_888 = lean_ctor_get(x_886, 0); lean_inc(x_888); lean_dec(x_886); -x_889 = lean_environment_main_module(x_888); +x_889 = l_Lean_Environment_mainModule(x_888); +lean_dec(x_888); x_890 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_891 = l_Lean_addMacroScope(x_889, x_890, x_846); lean_inc(x_2); @@ -31669,7 +31835,8 @@ lean_dec(x_847); x_902 = lean_ctor_get(x_900, 0); lean_inc(x_902); lean_dec(x_900); -x_903 = lean_environment_main_module(x_902); +x_903 = l_Lean_Environment_mainModule(x_902); +lean_dec(x_902); x_904 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_846); lean_inc(x_903); @@ -31743,7 +31910,8 @@ if (lean_is_exclusive(x_923)) { x_927 = lean_ctor_get(x_924, 0); lean_inc(x_927); lean_dec(x_924); -x_928 = lean_environment_main_module(x_927); +x_928 = l_Lean_Environment_mainModule(x_927); +lean_dec(x_927); x_929 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; x_930 = l_Lean_addMacroScope(x_928, x_929, x_846); lean_inc(x_2); @@ -32139,7 +32307,8 @@ x_42 = lean_ctor_get(x_40, 0); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); lean_dec(x_42); -x_44 = lean_environment_main_module(x_43); +x_44 = l_Lean_Environment_mainModule(x_43); +lean_dec(x_43); x_45 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_38); x_46 = lean_alloc_ctor(2, 2, 0); @@ -32276,7 +32445,8 @@ lean_dec(x_40); x_103 = lean_ctor_get(x_101, 0); lean_inc(x_103); lean_dec(x_101); -x_104 = lean_environment_main_module(x_103); +x_104 = l_Lean_Environment_mainModule(x_103); +lean_dec(x_103); x_105 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_38); x_106 = lean_alloc_ctor(2, 2, 0); @@ -32429,7 +32599,8 @@ x_168 = lean_ctor_get(x_166, 0); x_169 = lean_ctor_get(x_168, 0); lean_inc(x_169); lean_dec(x_168); -x_170 = lean_environment_main_module(x_169); +x_170 = l_Lean_Environment_mainModule(x_169); +lean_dec(x_169); x_171 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_164); x_172 = lean_alloc_ctor(2, 2, 0); @@ -32558,7 +32729,8 @@ lean_dec(x_166); x_224 = lean_ctor_get(x_222, 0); lean_inc(x_224); lean_dec(x_222); -x_225 = lean_environment_main_module(x_224); +x_225 = l_Lean_Environment_mainModule(x_224); +lean_dec(x_224); x_226 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_164); x_227 = lean_alloc_ctor(2, 2, 0); @@ -32703,7 +32875,8 @@ x_284 = lean_ctor_get(x_282, 0); x_285 = lean_ctor_get(x_284, 0); lean_inc(x_285); lean_dec(x_284); -x_286 = lean_environment_main_module(x_285); +x_286 = l_Lean_Environment_mainModule(x_285); +lean_dec(x_285); x_287 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_280); x_288 = lean_alloc_ctor(2, 2, 0); @@ -32867,7 +33040,8 @@ lean_dec(x_282); x_353 = lean_ctor_get(x_351, 0); lean_inc(x_353); lean_dec(x_351); -x_354 = lean_environment_main_module(x_353); +x_354 = l_Lean_Environment_mainModule(x_353); +lean_dec(x_353); x_355 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_280); x_356 = lean_alloc_ctor(2, 2, 0); @@ -33232,7 +33406,8 @@ x_20 = lean_ctor_get(x_17, 1); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -x_22 = lean_environment_main_module(x_21); +x_22 = l_Lean_Environment_mainModule(x_21); +lean_dec(x_21); x_23 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2; lean_inc(x_16); x_24 = l_Lean_addMacroScope(x_22, x_23, x_16); @@ -33297,7 +33472,8 @@ x_41 = lean_ctor_get(x_39, 0); x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); lean_dec(x_41); -x_43 = lean_environment_main_module(x_42); +x_43 = l_Lean_Environment_mainModule(x_42); +lean_dec(x_42); x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6; x_45 = l_Lean_addMacroScope(x_43, x_44, x_16); lean_inc(x_1); @@ -33337,7 +33513,8 @@ lean_dec(x_39); x_57 = lean_ctor_get(x_55, 0); lean_inc(x_57); lean_dec(x_55); -x_58 = lean_environment_main_module(x_57); +x_58 = l_Lean_Environment_mainModule(x_57); +lean_dec(x_57); x_59 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6; x_60 = l_Lean_addMacroScope(x_58, x_59, x_16); lean_inc(x_1); @@ -33485,7 +33662,8 @@ if (lean_is_exclusive(x_87)) { x_91 = lean_ctor_get(x_88, 0); lean_inc(x_91); lean_dec(x_88); -x_92 = lean_environment_main_module(x_91); +x_92 = l_Lean_Environment_mainModule(x_91); +lean_dec(x_91); x_93 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6; x_94 = l_Lean_addMacroScope(x_92, x_93, x_16); lean_inc(x_1); @@ -33599,7 +33777,8 @@ lean_dec(x_17); x_116 = lean_ctor_get(x_114, 0); lean_inc(x_116); lean_dec(x_114); -x_117 = lean_environment_main_module(x_116); +x_117 = l_Lean_Environment_mainModule(x_116); +lean_dec(x_116); x_118 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2; lean_inc(x_16); x_119 = l_Lean_addMacroScope(x_117, x_118, x_16); @@ -33677,7 +33856,8 @@ if (lean_is_exclusive(x_135)) { x_139 = lean_ctor_get(x_136, 0); lean_inc(x_139); lean_dec(x_136); -x_140 = lean_environment_main_module(x_139); +x_140 = l_Lean_Environment_mainModule(x_139); +lean_dec(x_139); x_141 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6; x_142 = l_Lean_addMacroScope(x_140, x_141, x_16); lean_inc(x_1); @@ -33809,7 +33989,8 @@ x_18 = lean_ctor_get(x_16, 0); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); lean_dec(x_18); -x_20 = lean_environment_main_module(x_19); +x_20 = l_Lean_Environment_mainModule(x_19); +lean_dec(x_19); x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_14); x_22 = lean_alloc_ctor(2, 2, 0); @@ -33905,7 +34086,8 @@ lean_dec(x_16); x_61 = lean_ctor_get(x_59, 0); lean_inc(x_61); lean_dec(x_59); -x_62 = lean_environment_main_module(x_61); +x_62 = l_Lean_Environment_mainModule(x_61); +lean_dec(x_61); x_63 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_14); x_64 = lean_alloc_ctor(2, 2, 0); @@ -34015,7 +34197,8 @@ x_17 = lean_ctor_get(x_15, 0); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); lean_dec(x_17); -x_19 = lean_environment_main_module(x_18); +x_19 = l_Lean_Environment_mainModule(x_18); +lean_dec(x_18); x_20 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_13); x_21 = lean_alloc_ctor(2, 2, 0); @@ -34074,7 +34257,8 @@ lean_dec(x_15); x_44 = lean_ctor_get(x_42, 0); lean_inc(x_44); lean_dec(x_42); -x_45 = lean_environment_main_module(x_44); +x_45 = l_Lean_Environment_mainModule(x_44); +lean_dec(x_44); x_46 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_13); x_47 = lean_alloc_ctor(2, 2, 0); @@ -36549,7 +36733,8 @@ x_25 = lean_ctor_get(x_22, 1); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -x_27 = lean_environment_main_module(x_26); +x_27 = l_Lean_Environment_mainModule(x_26); +lean_dec(x_26); x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__3; lean_inc(x_21); x_29 = l_Lean_addMacroScope(x_27, x_28, x_21); @@ -36584,7 +36769,8 @@ x_41 = lean_ctor_get(x_39, 0); x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); lean_dec(x_41); -x_43 = lean_environment_main_module(x_42); +x_43 = l_Lean_Environment_mainModule(x_42); +lean_dec(x_42); x_44 = l_Lean_addMacroScope(x_43, x_28, x_21); lean_inc(x_20); x_45 = lean_alloc_ctor(3, 4, 0); @@ -36628,7 +36814,8 @@ lean_dec(x_39); x_57 = lean_ctor_get(x_55, 0); lean_inc(x_57); lean_dec(x_55); -x_58 = lean_environment_main_module(x_57); +x_58 = l_Lean_Environment_mainModule(x_57); +lean_dec(x_57); x_59 = l_Lean_addMacroScope(x_58, x_28, x_21); lean_inc(x_20); x_60 = lean_alloc_ctor(3, 4, 0); @@ -36675,7 +36862,8 @@ lean_dec(x_22); x_73 = lean_ctor_get(x_71, 0); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); +x_74 = l_Lean_Environment_mainModule(x_73); +lean_dec(x_73); x_75 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__3; lean_inc(x_21); x_76 = l_Lean_addMacroScope(x_74, x_75, x_21); @@ -36717,7 +36905,8 @@ if (lean_is_exclusive(x_86)) { x_90 = lean_ctor_get(x_87, 0); lean_inc(x_90); lean_dec(x_87); -x_91 = lean_environment_main_module(x_90); +x_91 = l_Lean_Environment_mainModule(x_90); +lean_dec(x_90); x_92 = l_Lean_addMacroScope(x_91, x_75, x_21); lean_inc(x_20); x_93 = lean_alloc_ctor(3, 4, 0); @@ -36779,7 +36968,8 @@ x_112 = lean_ctor_get(x_109, 1); x_113 = lean_ctor_get(x_111, 0); lean_inc(x_113); lean_dec(x_111); -x_114 = lean_environment_main_module(x_113); +x_114 = l_Lean_Environment_mainModule(x_113); +lean_dec(x_113); x_115 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__3; lean_inc(x_108); lean_inc(x_114); @@ -36819,7 +37009,8 @@ x_131 = lean_ctor_get(x_129, 0); x_132 = lean_ctor_get(x_131, 0); lean_inc(x_132); lean_dec(x_131); -x_133 = lean_environment_main_module(x_132); +x_133 = l_Lean_Environment_mainModule(x_132); +lean_dec(x_132); x_134 = l_Lean_addMacroScope(x_133, x_115, x_108); lean_inc(x_107); x_135 = lean_alloc_ctor(3, 4, 0); @@ -36874,7 +37065,8 @@ lean_dec(x_129); x_153 = lean_ctor_get(x_151, 0); lean_inc(x_153); lean_dec(x_151); -x_154 = lean_environment_main_module(x_153); +x_154 = l_Lean_Environment_mainModule(x_153); +lean_dec(x_153); x_155 = l_Lean_addMacroScope(x_154, x_115, x_108); lean_inc(x_107); x_156 = lean_alloc_ctor(3, 4, 0); @@ -36932,7 +37124,8 @@ lean_dec(x_109); x_175 = lean_ctor_get(x_173, 0); lean_inc(x_175); lean_dec(x_173); -x_176 = lean_environment_main_module(x_175); +x_176 = l_Lean_Environment_mainModule(x_175); +lean_dec(x_175); x_177 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__3; lean_inc(x_108); lean_inc(x_176); @@ -36979,7 +37172,8 @@ if (lean_is_exclusive(x_191)) { x_195 = lean_ctor_get(x_192, 0); lean_inc(x_195); lean_dec(x_192); -x_196 = lean_environment_main_module(x_195); +x_196 = l_Lean_Environment_mainModule(x_195); +lean_dec(x_195); x_197 = l_Lean_addMacroScope(x_196, x_177, x_108); lean_inc(x_107); x_198 = lean_alloc_ctor(3, 4, 0); @@ -39312,7 +39506,8 @@ x_10 = lean_ctor_get(x_8, 0); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); -x_12 = lean_environment_main_module(x_11); +x_12 = l_Lean_Environment_mainModule(x_11); +lean_dec(x_11); x_13 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; x_14 = l_Lean_addMacroScope(x_12, x_13, x_7); x_15 = lean_box(0); @@ -39336,7 +39531,8 @@ lean_dec(x_8); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; x_23 = l_Lean_addMacroScope(x_21, x_22, x_7); x_24 = lean_box(0); @@ -40187,7 +40383,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 0); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_38); +x_39 = l_Lean_Environment_mainModule(x_38); +lean_dec(x_38); x_40 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_32); lean_ctor_set_tag(x_34, 2); @@ -40279,7 +40476,8 @@ lean_dec(x_34); x_77 = lean_ctor_get(x_75, 0); lean_inc(x_77); lean_dec(x_75); -x_78 = lean_environment_main_module(x_77); +x_78 = l_Lean_Environment_mainModule(x_77); +lean_dec(x_77); x_79 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_32); x_80 = lean_alloc_ctor(2, 2, 0); @@ -40392,7 +40590,8 @@ if (lean_is_exclusive(x_121)) { x_125 = lean_ctor_get(x_122, 0); lean_inc(x_125); lean_dec(x_122); -x_126 = lean_environment_main_module(x_125); +x_126 = l_Lean_Environment_mainModule(x_125); +lean_dec(x_125); x_127 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__19; lean_inc(x_119); if (lean_is_scalar(x_124)) { @@ -40907,7 +41106,8 @@ x_90 = lean_ctor_get(x_88, 0); x_91 = lean_ctor_get(x_90, 0); lean_inc(x_91); lean_dec(x_90); -x_92 = lean_environment_main_module(x_91); +x_92 = l_Lean_Environment_mainModule(x_91); +lean_dec(x_91); x_93 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_86); lean_ctor_set_tag(x_80, 2); @@ -40965,7 +41165,8 @@ lean_dec(x_88); x_113 = lean_ctor_get(x_111, 0); lean_inc(x_113); lean_dec(x_111); -x_114 = lean_environment_main_module(x_113); +x_114 = l_Lean_Environment_mainModule(x_113); +lean_dec(x_113); x_115 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_86); lean_ctor_set_tag(x_80, 2); @@ -41048,7 +41249,8 @@ if (lean_is_exclusive(x_140)) { x_144 = lean_ctor_get(x_141, 0); lean_inc(x_144); lean_dec(x_141); -x_145 = lean_environment_main_module(x_144); +x_145 = l_Lean_Environment_mainModule(x_144); +lean_dec(x_144); x_146 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_138); x_147 = lean_alloc_ctor(2, 2, 0); @@ -41217,7 +41419,8 @@ if (lean_is_exclusive(x_189)) { x_193 = lean_ctor_get(x_190, 0); lean_inc(x_193); lean_dec(x_190); -x_194 = lean_environment_main_module(x_193); +x_194 = l_Lean_Environment_mainModule(x_193); +lean_dec(x_193); x_195 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_187); if (lean_is_scalar(x_184)) { @@ -41402,7 +41605,8 @@ if (lean_is_exclusive(x_241)) { x_245 = lean_ctor_get(x_242, 0); lean_inc(x_245); lean_dec(x_242); -x_246 = lean_environment_main_module(x_245); +x_246 = l_Lean_Environment_mainModule(x_245); +lean_dec(x_245); x_247 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_239); if (lean_is_scalar(x_236)) { @@ -41779,7 +41983,8 @@ if (lean_is_exclusive(x_331)) { x_335 = lean_ctor_get(x_332, 0); lean_inc(x_335); lean_dec(x_332); -x_336 = lean_environment_main_module(x_335); +x_336 = l_Lean_Environment_mainModule(x_335); +lean_dec(x_335); x_337 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_329); if (lean_is_scalar(x_326)) { @@ -42183,7 +42388,8 @@ if (lean_is_exclusive(x_423)) { x_427 = lean_ctor_get(x_424, 0); lean_inc(x_427); lean_dec(x_424); -x_428 = lean_environment_main_module(x_427); +x_428 = l_Lean_Environment_mainModule(x_427); +lean_dec(x_427); x_429 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_421); if (lean_is_scalar(x_418)) { @@ -43104,7 +43310,8 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -x_14 = lean_environment_main_module(x_13); +x_14 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_15 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__6___lambda__1___closed__3; lean_inc(x_8); x_16 = l_Lean_addMacroScope(x_14, x_15, x_8); @@ -43125,7 +43332,8 @@ x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec(x_22); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l_Lean_addMacroScope(x_24, x_15, x_8); lean_inc(x_7); x_26 = lean_alloc_ctor(3, 4, 0); @@ -43154,7 +43362,8 @@ lean_dec(x_20); x_33 = lean_ctor_get(x_31, 0); lean_inc(x_33); lean_dec(x_31); -x_34 = lean_environment_main_module(x_33); +x_34 = l_Lean_Environment_mainModule(x_33); +lean_dec(x_33); x_35 = l_Lean_addMacroScope(x_34, x_15, x_8); lean_inc(x_7); x_36 = lean_alloc_ctor(3, 4, 0); @@ -43186,7 +43395,8 @@ lean_dec(x_9); x_44 = lean_ctor_get(x_42, 0); lean_inc(x_44); lean_dec(x_42); -x_45 = lean_environment_main_module(x_44); +x_45 = l_Lean_Environment_mainModule(x_44); +lean_dec(x_44); x_46 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__6___lambda__1___closed__3; lean_inc(x_8); x_47 = l_Lean_addMacroScope(x_45, x_46, x_8); @@ -43214,7 +43424,8 @@ if (lean_is_exclusive(x_51)) { x_55 = lean_ctor_get(x_52, 0); lean_inc(x_55); lean_dec(x_52); -x_56 = lean_environment_main_module(x_55); +x_56 = l_Lean_Environment_mainModule(x_55); +lean_dec(x_55); x_57 = l_Lean_addMacroScope(x_56, x_46, x_8); lean_inc(x_7); x_58 = lean_alloc_ctor(3, 4, 0); diff --git a/stage0/stdlib/Lean/Elab/Quotation/Precheck.c b/stage0/stdlib/Lean/Elab/Quotation/Precheck.c index cb02223ea10d..34cd97cb134f 100644 --- a/stage0/stdlib/Lean/Elab/Quotation/Precheck.c +++ b/stage0/stdlib/Lean/Elab/Quotation/Precheck.c @@ -248,7 +248,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_precheckApp__1___clo static lean_object* l_Lean_Elab_Term_Quotation_runPrecheck___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_precheckBinrel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_precheckExplicit___closed__2; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Elab_Term_Quotation_precheckChoice___spec__4___boxed(lean_object*, lean_object*, lean_object*); @@ -2070,7 +2070,8 @@ x_29 = lean_ctor_get(x_26, 1); x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_13); +x_31 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_32 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_32, 0, x_25); lean_ctor_set(x_32, 1, x_31); @@ -2248,7 +2249,8 @@ lean_dec(x_26); x_80 = lean_ctor_get(x_78, 1); lean_inc(x_80); lean_dec(x_78); -x_81 = lean_environment_main_module(x_13); +x_81 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_82 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_82, 0, x_25); lean_ctor_set(x_82, 1, x_81); diff --git a/stage0/stdlib/Lean/Elab/StructInst.c b/stage0/stdlib/Lean/Elab/StructInst.c index 1ec3475a0bf6..945f7eedb69f 100644 --- a/stage0/stdlib/Lean/Elab/StructInst.c +++ b/stage0/stdlib/Lean/Elab/StructInst.c @@ -649,7 +649,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Stru LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_elabStructInst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStructInstField___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Lean_Meta_synthInstance_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault_loop___closed__1; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStructInstField___lambda__2___closed__4; @@ -3591,7 +3591,8 @@ x_19 = lean_ctor_get(x_16, 1); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__3; lean_inc(x_15); x_23 = l_Lean_addMacroScope(x_21, x_22, x_15); @@ -3624,7 +3625,8 @@ x_33 = lean_ctor_get(x_31, 0); x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); lean_dec(x_33); -x_35 = lean_environment_main_module(x_34); +x_35 = l_Lean_Environment_mainModule(x_34); +lean_dec(x_34); x_36 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__4; lean_inc(x_14); lean_ctor_set_tag(x_16, 2); @@ -3677,7 +3679,8 @@ lean_dec(x_31); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__4; lean_inc(x_14); lean_ctor_set_tag(x_16, 2); @@ -3761,7 +3764,8 @@ lean_dec(x_16); x_79 = lean_ctor_get(x_77, 0); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_79); +x_80 = l_Lean_Environment_mainModule(x_79); +lean_dec(x_79); x_81 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__3; lean_inc(x_15); x_82 = l_Lean_addMacroScope(x_80, x_81, x_15); @@ -3801,7 +3805,8 @@ if (lean_is_exclusive(x_90)) { x_94 = lean_ctor_get(x_91, 0); lean_inc(x_94); lean_dec(x_91); -x_95 = lean_environment_main_module(x_94); +x_95 = l_Lean_Environment_mainModule(x_94); +lean_dec(x_94); x_96 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__4; lean_inc(x_14); x_97 = lean_alloc_ctor(2, 2, 0); @@ -6374,7 +6379,8 @@ lean_dec(x_30); x_35 = lean_ctor_get(x_32, 0); lean_inc(x_35); lean_dec(x_32); -x_36 = lean_environment_main_module(x_35); +x_36 = l_Lean_Environment_mainModule(x_35); +lean_dec(x_35); x_37 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__5; lean_inc(x_27); lean_ctor_set_tag(x_29, 2); @@ -6572,7 +6578,8 @@ lean_dec(x_30); x_110 = lean_ctor_get(x_107, 0); lean_inc(x_110); lean_dec(x_107); -x_111 = lean_environment_main_module(x_110); +x_111 = l_Lean_Environment_mainModule(x_110); +lean_dec(x_110); x_112 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__5; lean_inc(x_27); x_113 = lean_alloc_ctor(2, 2, 0); @@ -6844,7 +6851,8 @@ x_25 = lean_ctor_get(x_22, 1); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -x_27 = lean_environment_main_module(x_26); +x_27 = l_Lean_Environment_mainModule(x_26); +lean_dec(x_26); x_28 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__3; x_29 = l_Lean_addMacroScope(x_27, x_28, x_21); x_30 = lean_box(0); @@ -7072,7 +7080,8 @@ lean_dec(x_22); x_117 = lean_ctor_get(x_115, 0); lean_inc(x_117); lean_dec(x_115); -x_118 = lean_environment_main_module(x_117); +x_118 = l_Lean_Environment_mainModule(x_117); +lean_dec(x_117); x_119 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__3; x_120 = l_Lean_addMacroScope(x_118, x_119, x_21); x_121 = lean_box(0); @@ -7276,7 +7285,8 @@ x_194 = lean_ctor_get(x_191, 1); x_195 = lean_ctor_get(x_193, 0); lean_inc(x_195); lean_dec(x_193); -x_196 = lean_environment_main_module(x_195); +x_196 = l_Lean_Environment_mainModule(x_195); +lean_dec(x_195); x_197 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__5; lean_inc(x_189); lean_ctor_set_tag(x_191, 2); @@ -7470,7 +7480,8 @@ lean_dec(x_191); x_274 = lean_ctor_get(x_272, 0); lean_inc(x_274); lean_dec(x_272); -x_275 = lean_environment_main_module(x_274); +x_275 = l_Lean_Environment_mainModule(x_274); +lean_dec(x_274); x_276 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__5; lean_inc(x_189); x_277 = lean_alloc_ctor(2, 2, 0); @@ -28704,7 +28715,6 @@ x_20 = lean_ctor_get(x_17, 1); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_11); lean_inc(x_21); x_22 = l_Lean_getStructureCtor(x_21, x_11); x_23 = lean_ctor_get(x_22, 0); @@ -28712,7 +28722,6 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); lean_dec(x_23); -lean_inc(x_21); x_25 = l_Lean_isPrivateNameFromImportedModule(x_21, x_24); if (x_25 == 0) { @@ -28777,7 +28786,6 @@ lean_dec(x_17); x_39 = lean_ctor_get(x_37, 0); lean_inc(x_39); lean_dec(x_37); -lean_inc(x_11); lean_inc(x_39); x_40 = l_Lean_getStructureCtor(x_39, x_11); x_41 = lean_ctor_get(x_40, 0); @@ -28785,7 +28793,6 @@ lean_inc(x_41); x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); lean_dec(x_41); -lean_inc(x_39); x_43 = l_Lean_isPrivateNameFromImportedModule(x_39, x_42); if (x_43 == 0) { @@ -28904,7 +28911,6 @@ if (lean_is_exclusive(x_72)) { x_76 = lean_ctor_get(x_73, 0); lean_inc(x_76); lean_dec(x_73); -lean_inc(x_11); lean_inc(x_76); x_77 = l_Lean_getStructureCtor(x_76, x_11); x_78 = lean_ctor_get(x_77, 0); @@ -28912,7 +28918,6 @@ lean_inc(x_78); x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); lean_dec(x_78); -lean_inc(x_76); x_80 = l_Lean_isPrivateNameFromImportedModule(x_76, x_79); if (x_80 == 0) { @@ -38859,7 +38864,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -39041,7 +39047,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); diff --git a/stage0/stdlib/Lean/Elab/Structure.c b/stage0/stdlib/Lean/Elab/Structure.c index ae938f66b0e8..c0f1cd251a5f 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -916,7 +916,7 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkToPar uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFieldTypeValue___lambda__1___boxed(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_elabStructureCommand___elambda__1___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructureCommand___elambda__1___lambda__4___boxed__const__1; @@ -924,6 +924,7 @@ lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta lean_object* l_Lean_MessageData_ofLevel(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getNestedProjectionArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedStructFieldKind; lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__3; @@ -1243,7 +1244,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Structu LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructureCommand___elambda__1___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__10; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___closed__1; -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); uint8_t lean_is_class(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Lean_throwKernelException___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3459,7 +3459,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -3641,7 +3642,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); @@ -8129,7 +8131,8 @@ x_87 = lean_ctor_get(x_84, 1); x_88 = lean_ctor_get(x_86, 0); lean_inc(x_88); lean_dec(x_86); -x_89 = lean_environment_main_module(x_88); +x_89 = l_Lean_Environment_mainModule(x_88); +lean_dec(x_88); x_90 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__13; x_91 = l_Lean_addMacroScope(x_89, x_90, x_83); x_92 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__12; @@ -8172,7 +8175,8 @@ lean_dec(x_84); x_104 = lean_ctor_get(x_102, 0); lean_inc(x_104); lean_dec(x_102); -x_105 = lean_environment_main_module(x_104); +x_105 = l_Lean_Environment_mainModule(x_104); +lean_dec(x_104); x_106 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__13; x_107 = l_Lean_addMacroScope(x_105, x_106, x_83); x_108 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__12; @@ -8259,7 +8263,8 @@ if (lean_is_exclusive(x_133)) { x_137 = lean_ctor_get(x_134, 0); lean_inc(x_137); lean_dec(x_134); -x_138 = lean_environment_main_module(x_137); +x_138 = l_Lean_Environment_mainModule(x_137); +lean_dec(x_137); x_139 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__13; x_140 = l_Lean_addMacroScope(x_138, x_139, x_132); x_141 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__12; @@ -8427,7 +8432,8 @@ if (lean_is_exclusive(x_181)) { x_185 = lean_ctor_get(x_182, 0); lean_inc(x_185); lean_dec(x_182); -x_186 = lean_environment_main_module(x_185); +x_186 = l_Lean_Environment_mainModule(x_185); +lean_dec(x_185); x_187 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__13; x_188 = l_Lean_addMacroScope(x_186, x_187, x_180); x_189 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__12; @@ -8705,7 +8711,8 @@ if (lean_is_exclusive(x_256)) { x_260 = lean_ctor_get(x_257, 0); lean_inc(x_260); lean_dec(x_257); -x_261 = lean_environment_main_module(x_260); +x_261 = l_Lean_Environment_mainModule(x_260); +lean_dec(x_260); x_262 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__13; x_263 = l_Lean_addMacroScope(x_261, x_262, x_255); x_264 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___closed__12; @@ -15271,7 +15278,6 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -lean_inc(x_15); lean_inc(x_13); x_17 = l_Lean_getStructureCtor(x_13, x_15); x_18 = lean_ctor_get(x_17, 0); @@ -24872,7 +24878,7 @@ LEAN_EXPORT lean_object* l_Lean_throwKernelException___at___private_Lean_Elab_St lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = lean_ctor_get(x_6, 2); lean_inc(x_9); -x_10 = l_Lean_KernelException_toMessageData(x_1, x_9); +x_10 = l_Lean_Kernel_Exception_toMessageData(x_1, x_9); x_11 = l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__4(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_6); return x_11; @@ -32283,7 +32289,6 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_11); lean_inc(x_2); x_13 = l_Lean_getStructureCtor(x_2, x_11); x_14 = lean_ctor_get(x_13, 0); diff --git a/stage0/stdlib/Lean/Elab/Syntax.c b/stage0/stdlib/Lean/Elab/Syntax.c index 9482ecd417c9..045231ece633 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -592,7 +592,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_strLitToPattern(lean_object*, lean_ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkRuleKind___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Substring_takeRightWhileAux___at_Substring_trimRight___spec__1(lean_object*, lean_object*, lean_object*); @@ -1169,7 +1169,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_29); +x_30 = l_Lean_Environment_mainModule(x_29); +lean_dec(x_29); x_31 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__10; x_32 = l_Lean_addMacroScope(x_30, x_31, x_24); x_33 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; @@ -1219,7 +1220,8 @@ lean_dec(x_25); x_51 = lean_ctor_get(x_49, 0); lean_inc(x_51); lean_dec(x_49); -x_52 = lean_environment_main_module(x_51); +x_52 = l_Lean_Environment_mainModule(x_51); +lean_dec(x_51); x_53 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__10; x_54 = l_Lean_addMacroScope(x_52, x_53, x_24); x_55 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; @@ -1295,7 +1297,8 @@ if (lean_is_exclusive(x_80)) { x_84 = lean_ctor_get(x_81, 0); lean_inc(x_84); lean_dec(x_81); -x_85 = lean_environment_main_module(x_84); +x_85 = l_Lean_Environment_mainModule(x_84); +lean_dec(x_84); x_86 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__10; x_87 = l_Lean_addMacroScope(x_85, x_86, x_79); x_88 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; @@ -3166,7 +3169,8 @@ x_30 = lean_ctor_get(x_27, 1); x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_environment_main_module(x_14); +x_32 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_33 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_33, 0, x_26); lean_ctor_set(x_33, 1, x_32); @@ -3344,7 +3348,8 @@ lean_dec(x_27); x_81 = lean_ctor_get(x_79, 1); lean_inc(x_81); lean_dec(x_79); -x_82 = lean_environment_main_module(x_14); +x_82 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_83 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_83, 0, x_26); lean_ctor_set(x_83, 1, x_82); @@ -5760,7 +5765,8 @@ x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec(x_22); -x_24 = lean_environment_main_module(x_23); +x_24 = l_Lean_Environment_mainModule(x_23); +lean_dec(x_23); x_25 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; lean_inc(x_18); x_26 = lean_alloc_ctor(2, 2, 0); @@ -5839,7 +5845,8 @@ lean_dec(x_20); x_60 = lean_ctor_get(x_58, 0); lean_inc(x_60); lean_dec(x_58); -x_61 = lean_environment_main_module(x_60); +x_61 = l_Lean_Environment_mainModule(x_60); +lean_dec(x_60); x_62 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; lean_inc(x_18); x_63 = lean_alloc_ctor(2, 2, 0); @@ -6392,7 +6399,8 @@ x_65 = lean_ctor_get(x_63, 0); x_66 = lean_ctor_get(x_65, 0); lean_inc(x_66); lean_dec(x_65); -x_67 = lean_environment_main_module(x_66); +x_67 = l_Lean_Environment_mainModule(x_66); +lean_dec(x_66); x_68 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__13; lean_inc(x_62); lean_inc(x_67); @@ -6440,7 +6448,8 @@ lean_dec(x_63); x_88 = lean_ctor_get(x_86, 0); lean_inc(x_88); lean_dec(x_86); -x_89 = lean_environment_main_module(x_88); +x_89 = l_Lean_Environment_mainModule(x_88); +lean_dec(x_88); x_90 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__13; lean_inc(x_62); lean_inc(x_89); @@ -6509,7 +6518,8 @@ x_19 = lean_ctor_get(x_17, 0); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); lean_dec(x_19); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__4; x_23 = l_Lean_addMacroScope(x_21, x_22, x_16); x_24 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__2; @@ -6545,7 +6555,8 @@ lean_dec(x_17); x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_environment_main_module(x_37); +x_38 = l_Lean_Environment_mainModule(x_37); +lean_dec(x_37); x_39 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__4; x_40 = l_Lean_addMacroScope(x_38, x_39, x_16); x_41 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__2; @@ -6821,7 +6832,8 @@ x_26 = lean_ctor_get(x_23, 0); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); lean_dec(x_26); -x_28 = lean_environment_main_module(x_27); +x_28 = l_Lean_Environment_mainModule(x_27); +lean_dec(x_27); x_29 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__3; x_30 = l_Lean_addMacroScope(x_28, x_29, x_22); x_31 = lean_box(0); @@ -6898,7 +6910,8 @@ lean_dec(x_23); x_62 = lean_ctor_get(x_60, 0); lean_inc(x_62); lean_dec(x_60); -x_63 = lean_environment_main_module(x_62); +x_63 = l_Lean_Environment_mainModule(x_62); +lean_dec(x_62); x_64 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__3; x_65 = l_Lean_addMacroScope(x_63, x_64, x_22); x_66 = lean_box(0); @@ -7410,7 +7423,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -7592,7 +7606,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); @@ -8768,7 +8783,8 @@ x_24 = lean_ctor_get(x_21, 0); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); lean_dec(x_24); -x_26 = lean_environment_main_module(x_25); +x_26 = l_Lean_Environment_mainModule(x_25); +lean_dec(x_25); x_27 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; lean_inc(x_19); x_28 = lean_alloc_ctor(2, 2, 0); @@ -8854,7 +8870,8 @@ lean_dec(x_21); x_64 = lean_ctor_get(x_62, 0); lean_inc(x_64); lean_dec(x_62); -x_65 = lean_environment_main_module(x_64); +x_65 = l_Lean_Environment_mainModule(x_64); +lean_dec(x_64); x_66 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; lean_inc(x_19); x_67 = lean_alloc_ctor(2, 2, 0); @@ -9057,7 +9074,8 @@ lean_dec(x_39); x_42 = lean_ctor_get(x_40, 0); lean_inc(x_42); lean_dec(x_40); -x_43 = lean_environment_main_module(x_42); +x_43 = l_Lean_Environment_mainModule(x_42); +lean_dec(x_42); x_44 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__4; x_45 = l_Lean_addMacroScope(x_43, x_44, x_38); x_46 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__2; @@ -9240,7 +9258,8 @@ x_24 = lean_ctor_get(x_21, 0); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); lean_dec(x_24); -x_26 = lean_environment_main_module(x_25); +x_26 = l_Lean_Environment_mainModule(x_25); +lean_dec(x_25); x_27 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; lean_inc(x_19); x_28 = lean_alloc_ctor(2, 2, 0); @@ -9326,7 +9345,8 @@ lean_dec(x_21); x_64 = lean_ctor_get(x_62, 0); lean_inc(x_64); lean_dec(x_62); -x_65 = lean_environment_main_module(x_64); +x_65 = l_Lean_Environment_mainModule(x_64); +lean_dec(x_64); x_66 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; lean_inc(x_19); x_67 = lean_alloc_ctor(2, 2, 0); @@ -9529,7 +9549,8 @@ lean_dec(x_39); x_42 = lean_ctor_get(x_40, 0); lean_inc(x_42); lean_dec(x_40); -x_43 = lean_environment_main_module(x_42); +x_43 = l_Lean_Environment_mainModule(x_42); +lean_dec(x_42); x_44 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__4; x_45 = l_Lean_addMacroScope(x_43, x_44, x_38); x_46 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__2; @@ -9927,7 +9948,8 @@ x_19 = lean_ctor_get(x_17, 0); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); lean_dec(x_19); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__6___lambda__1___closed__4; x_23 = l_Lean_addMacroScope(x_21, x_22, x_16); x_24 = lean_box(0); @@ -10013,7 +10035,8 @@ lean_dec(x_17); x_60 = lean_ctor_get(x_58, 0); lean_inc(x_60); lean_dec(x_58); -x_61 = lean_environment_main_module(x_60); +x_61 = l_Lean_Environment_mainModule(x_60); +lean_dec(x_60); x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__6___lambda__1___closed__4; x_63 = l_Lean_addMacroScope(x_61, x_62, x_16); x_64 = lean_box(0); @@ -10583,7 +10606,8 @@ x_43 = lean_ctor_get(x_40, 1); x_44 = lean_ctor_get(x_42, 0); lean_inc(x_44); lean_dec(x_42); -x_45 = lean_environment_main_module(x_44); +x_45 = l_Lean_Environment_mainModule(x_44); +lean_dec(x_44); x_46 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__10; x_47 = l_Lean_addMacroScope(x_45, x_46, x_39); x_48 = lean_box(0); @@ -10669,7 +10693,8 @@ lean_dec(x_40); x_76 = lean_ctor_get(x_74, 0); lean_inc(x_76); lean_dec(x_74); -x_77 = lean_environment_main_module(x_76); +x_77 = l_Lean_Environment_mainModule(x_76); +lean_dec(x_76); x_78 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__10; x_79 = l_Lean_addMacroScope(x_77, x_78, x_39); x_80 = lean_box(0); @@ -10828,7 +10853,8 @@ x_130 = lean_ctor_get(x_127, 1); x_131 = lean_ctor_get(x_129, 0); lean_inc(x_131); lean_dec(x_129); -x_132 = lean_environment_main_module(x_131); +x_132 = l_Lean_Environment_mainModule(x_131); +lean_dec(x_131); x_133 = l_Lean_Elab_Term_ensureUnaryOutput___closed__7; x_134 = l_Lean_addMacroScope(x_132, x_133, x_126); x_135 = lean_box(0); @@ -10914,7 +10940,8 @@ lean_dec(x_127); x_163 = lean_ctor_get(x_161, 0); lean_inc(x_163); lean_dec(x_161); -x_164 = lean_environment_main_module(x_163); +x_164 = l_Lean_Environment_mainModule(x_163); +lean_dec(x_163); x_165 = l_Lean_Elab_Term_ensureUnaryOutput___closed__7; x_166 = l_Lean_addMacroScope(x_164, x_165, x_126); x_167 = lean_box(0); @@ -11071,7 +11098,8 @@ x_216 = lean_ctor_get(x_213, 1); x_217 = lean_ctor_get(x_215, 0); lean_inc(x_217); lean_dec(x_215); -x_218 = lean_environment_main_module(x_217); +x_218 = l_Lean_Environment_mainModule(x_217); +lean_dec(x_217); x_219 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2___closed__8; x_220 = l_Lean_addMacroScope(x_218, x_219, x_212); x_221 = lean_box(0); @@ -11157,7 +11185,8 @@ lean_dec(x_213); x_249 = lean_ctor_get(x_247, 0); lean_inc(x_249); lean_dec(x_247); -x_250 = lean_environment_main_module(x_249); +x_250 = l_Lean_Environment_mainModule(x_249); +lean_dec(x_249); x_251 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2___closed__8; x_252 = l_Lean_addMacroScope(x_250, x_251, x_212); x_253 = lean_box(0); @@ -11565,7 +11594,8 @@ x_21 = lean_ctor_get(x_19, 0); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec(x_21); -x_23 = lean_environment_main_module(x_22); +x_23 = l_Lean_Environment_mainModule(x_22); +lean_dec(x_22); x_24 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___lambda__1___closed__4; x_25 = l_Lean_addMacroScope(x_23, x_24, x_18); x_26 = lean_box(0); @@ -11642,7 +11672,8 @@ lean_dec(x_19); x_57 = lean_ctor_get(x_55, 0); lean_inc(x_57); lean_dec(x_55); -x_58 = lean_environment_main_module(x_57); +x_58 = l_Lean_Environment_mainModule(x_57); +lean_dec(x_57); x_59 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___lambda__1___closed__4; x_60 = l_Lean_addMacroScope(x_58, x_59, x_18); x_61 = lean_box(0); @@ -18989,7 +19020,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 3); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_8); +x_39 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_40 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_40, 0, x_22); lean_ctor_set(x_40, 1, x_39); @@ -19170,7 +19202,8 @@ lean_dec(x_34); x_89 = lean_ctor_get(x_87, 3); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_8); +x_90 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_91 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_91, 0, x_22); lean_ctor_set(x_91, 1, x_90); @@ -19505,7 +19538,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 3); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_8); +x_39 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_40 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_40, 0, x_22); lean_ctor_set(x_40, 1, x_39); @@ -19686,7 +19720,8 @@ lean_dec(x_34); x_89 = lean_ctor_get(x_87, 3); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_8); +x_90 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_91 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_91, 0, x_22); lean_ctor_set(x_91, 1, x_90); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Basic.c b/stage0/stdlib/Lean/Elab/Tactic/Basic.c index 1803dd4138a8..d7b19bd76c3f 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Basic.c @@ -435,7 +435,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_throwNoGoalsToBeSolved___rarg___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instMonadTacticM___closed__2; static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8563____closed__11; @@ -4488,7 +4488,8 @@ x_5 = lean_ctor_get(x_3, 0); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); lean_dec(x_5); -x_7 = lean_environment_main_module(x_6); +x_7 = l_Lean_Environment_mainModule(x_6); +lean_dec(x_6); lean_ctor_set(x_3, 0, x_7); return x_3; } @@ -4503,7 +4504,8 @@ lean_dec(x_3); x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); -x_11 = lean_environment_main_module(x_10); +x_11 = l_Lean_Environment_mainModule(x_10); +lean_dec(x_10); x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_9); @@ -10029,7 +10031,8 @@ x_30 = lean_ctor_get(x_27, 1); x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_environment_main_module(x_14); +x_32 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_33 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_33, 0, x_26); lean_ctor_set(x_33, 1, x_32); @@ -10207,7 +10210,8 @@ lean_dec(x_27); x_81 = lean_ctor_get(x_79, 1); lean_inc(x_81); lean_dec(x_79); -x_82 = lean_environment_main_module(x_14); +x_82 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_83 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_83, 0, x_26); lean_ctor_set(x_83, 1, x_82); diff --git a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c index f8a296c97bb7..970fef33be14 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c @@ -275,7 +275,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalClear___lambda__1(lean_object*, static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRotateLeft_declRange__1___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRunTac_declRange__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRefl_declRange__1___closed__6; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntro_declRange__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDbgTrace___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -416,7 +416,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_BuiltinTactic_0__Lean_Elab static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalTraceState__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalOpen___spec__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_evalAllGoals___spec__1___boxed(lean_object**); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalChoice_declRange__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRange__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSepTactics_goOdd___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*); @@ -432,7 +431,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTacticSeqBracketed___lambda__1(l static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFail__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSetOption___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntro_declRange__1(lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Elab_Tactic_renameInaccessibles___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalRefl(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalContradiction__1___closed__4; @@ -655,6 +653,7 @@ lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRunTac_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFirst_declRange__1___closed__1; +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTraceMessage___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAssumption___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented_declRange__1___closed__3; @@ -1109,7 +1108,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalLeft_declRange__1___clos static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntroMatch__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRunTac_declRange__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSepTactics_goEven___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalContradiction_declRange__1___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalOpen_declRange__1___closed__5; @@ -1289,7 +1288,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRenameInaccessibles_decl lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_evalOpen___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSeq1_declRange__1___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_renameInaccessibles___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__6___closed__3; @@ -1308,7 +1307,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_renameInaccessibles___lambda__3(lean LEAN_EXPORT lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Tactic_evalOpen___spec__34___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalContradiction___boxed(lean_object*); lean_object* l_Lean_Elab_Term_evalTerm___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSeq1__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSkip_declRange__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRunTac__1___closed__3; @@ -1406,6 +1405,7 @@ lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSubstEqs_declRange__1___closed__2; lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState__1___closed__6; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalTraceState_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFailIfSuccess_declRange__1___closed__7; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_evalOpen___spec__20(lean_object*, lean_object*, lean_object*); @@ -15298,8 +15298,7 @@ x_15 = lean_ctor_get(x_12, 1); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -lean_inc(x_1); -x_17 = lean_environment_find(x_16, x_1); +x_17 = l_Lean_Environment_find_x3f(x_16, x_1); if (lean_obj_tag(x_17) == 0) { uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -15339,8 +15338,7 @@ lean_dec(x_12); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -lean_inc(x_1); -x_29 = lean_environment_find(x_28, x_1); +x_29 = l_Lean_Environment_find_x3f(x_28, x_1); if (lean_obj_tag(x_29) == 0) { uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; @@ -22693,7 +22691,7 @@ lean_object* x_29; uint8_t x_30; uint8_t x_31; x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -x_30 = l_Lean_Kernel_isDiagnosticsEnabled(x_29); +x_30 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_29); lean_dec(x_29); if (x_30 == 0) { @@ -22747,7 +22745,7 @@ lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean x_36 = lean_ctor_get(x_33, 0); x_37 = lean_ctor_get(x_33, 4); lean_dec(x_37); -x_38 = l_Lean_Kernel_enableDiag(x_36, x_24); +x_38 = l_Lean_Kernel_Environment_enableDiag(x_36, x_24); x_39 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__2___closed__3; lean_ctor_set(x_33, 4, x_39); lean_ctor_set(x_33, 0, x_38); @@ -22777,7 +22775,7 @@ lean_inc(x_46); lean_inc(x_45); lean_inc(x_44); lean_dec(x_33); -x_51 = l_Lean_Kernel_enableDiag(x_44, x_24); +x_51 = l_Lean_Kernel_Environment_enableDiag(x_44, x_24); x_52 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__2___closed__3; x_53 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_53, 0, x_51); @@ -28943,213 +28941,436 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_getMainGoal(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +lean_object* x_12; +x_12 = l_Lean_Elab_Tactic_getMainGoal(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -x_14 = l_Lean_MVarId_intro(x_12, x_1, x_6, x_7, x_8, x_9, x_13); -if (lean_obj_tag(x_14) == 0) +x_15 = l_Lean_MVarId_intro(x_13, x_2, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_15, 0); -x_19 = lean_ctor_get(x_15, 1); -x_20 = lean_box(0); -lean_ctor_set_tag(x_15, 1); -lean_ctor_set(x_15, 1, x_20); -lean_ctor_set(x_15, 0, x_19); -x_21 = l_Lean_Elab_Tactic_replaceMainGoal(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +x_21 = lean_box(0); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 1, x_21); +lean_ctor_set(x_16, 0, x_20); +x_22 = l_Lean_Elab_Tactic_replaceMainGoal(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -if (lean_obj_tag(x_21) == 0) +if (lean_obj_tag(x_22) == 0) { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_18); -return x_21; +lean_object* x_24; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +lean_ctor_set(x_22, 0, x_19); +return x_22; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_18); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_19); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -uint8_t x_26; -lean_dec(x_18); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) +uint8_t x_27; +lean_dec(x_19); +x_27 = !lean_is_exclusive(x_22); +if (x_27 == 0) { -return x_21; +return x_22; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_22, 0); +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_21); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_dec(x_22); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_30 = lean_ctor_get(x_15, 0); -x_31 = lean_ctor_get(x_15, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_16, 0); +x_32 = lean_ctor_get(x_16, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_15); -x_32 = lean_box(0); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Elab_Tactic_replaceMainGoal(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +lean_dec(x_16); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_Elab_Tactic_replaceMainGoal(x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -if (lean_obj_tag(x_34) == 0) +if (lean_obj_tag(x_35) == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_36 = x_34; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_37 = x_35; } else { - lean_dec_ref(x_34); - x_36 = lean_box(0); + lean_dec_ref(x_35); + x_37 = lean_box(0); } -if (lean_is_scalar(x_36)) { - x_37 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_37)) { + x_38 = lean_alloc_ctor(0, 2, 0); } else { - x_37 = x_36; + x_38 = x_37; } -lean_ctor_set(x_37, 0, x_30); -lean_ctor_set(x_37, 1, x_35); -return x_37; +lean_ctor_set(x_38, 0, x_31); +lean_ctor_set(x_38, 1, x_36); +return x_38; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -lean_dec(x_30); -x_38 = lean_ctor_get(x_34, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_34, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_31); +x_39 = lean_ctor_get(x_35, 0); lean_inc(x_39); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_40 = x_34; +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_41 = x_35; } else { - lean_dec_ref(x_34); - x_40 = lean_box(0); + lean_dec_ref(x_35); + x_41 = lean_box(0); } -if (lean_is_scalar(x_40)) { - x_41 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_41)) { + x_42 = lean_alloc_ctor(1, 2, 0); } else { - x_41 = x_40; + x_42 = x_41; } -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_40); +return x_42; } } } else { -uint8_t x_42; +uint8_t x_43; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -x_42 = !lean_is_exclusive(x_14); -if (x_42 == 0) +x_43 = !lean_is_exclusive(x_15); +if (x_43 == 0) { -return x_14; +return x_15; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_14, 0); -x_44 = lean_ctor_get(x_14, 1); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_15, 0); +x_45 = lean_ctor_get(x_15, 1); +lean_inc(x_45); lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_14); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; +lean_dec(x_15); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } else { -uint8_t x_46; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_47 = lean_ctor_get(x_12, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_12, 1); +lean_inc(x_48); +lean_dec(x_12); +x_49 = lean_ctor_get(x_1, 0); +x_50 = lean_ctor_get(x_9, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_9, 1); +lean_inc(x_51); +x_52 = lean_ctor_get(x_9, 2); +lean_inc(x_52); +x_53 = lean_ctor_get(x_9, 3); +lean_inc(x_53); +x_54 = lean_ctor_get(x_9, 4); +lean_inc(x_54); +x_55 = lean_ctor_get(x_9, 5); +lean_inc(x_55); +x_56 = lean_ctor_get(x_9, 6); +lean_inc(x_56); +x_57 = lean_ctor_get(x_9, 7); +lean_inc(x_57); +x_58 = lean_ctor_get(x_9, 8); +lean_inc(x_58); +x_59 = lean_ctor_get(x_9, 9); +lean_inc(x_59); +x_60 = lean_ctor_get(x_9, 10); +lean_inc(x_60); +x_61 = lean_ctor_get_uint8(x_9, sizeof(void*)*12); +x_62 = lean_ctor_get(x_9, 11); +lean_inc(x_62); +x_63 = lean_ctor_get_uint8(x_9, sizeof(void*)*12 + 1); +x_64 = l_Lean_replaceRef(x_49, x_55); +lean_dec(x_55); +x_65 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_65, 0, x_50); +lean_ctor_set(x_65, 1, x_51); +lean_ctor_set(x_65, 2, x_52); +lean_ctor_set(x_65, 3, x_53); +lean_ctor_set(x_65, 4, x_54); +lean_ctor_set(x_65, 5, x_64); +lean_ctor_set(x_65, 6, x_56); +lean_ctor_set(x_65, 7, x_57); +lean_ctor_set(x_65, 8, x_58); +lean_ctor_set(x_65, 9, x_59); +lean_ctor_set(x_65, 10, x_60); +lean_ctor_set(x_65, 11, x_62); +lean_ctor_set_uint8(x_65, sizeof(void*)*12, x_61); +lean_ctor_set_uint8(x_65, sizeof(void*)*12 + 1, x_63); +lean_inc(x_10); +lean_inc(x_8); +lean_inc(x_7); +x_66 = l_Lean_MVarId_intro(x_47, x_2, x_7, x_8, x_65, x_10, x_48); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = !lean_is_exclusive(x_67); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_67, 0); +x_71 = lean_ctor_get(x_67, 1); +x_72 = lean_box(0); +lean_ctor_set_tag(x_67, 1); +lean_ctor_set(x_67, 1, x_72); +lean_ctor_set(x_67, 0, x_71); +x_73 = l_Lean_Elab_Tactic_replaceMainGoal(x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_46 = !lean_is_exclusive(x_11); -if (x_46 == 0) +if (lean_obj_tag(x_73) == 0) { -return x_11; +uint8_t x_74; +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) +{ +lean_object* x_75; +x_75 = lean_ctor_get(x_73, 0); +lean_dec(x_75); +lean_ctor_set(x_73, 0, x_70); +return x_73; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_11, 0); -x_48 = lean_ctor_get(x_11, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_11); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_70); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +} +else +{ +uint8_t x_78; +lean_dec(x_70); +x_78 = !lean_is_exclusive(x_73); +if (x_78 == 0) +{ +return x_73; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_73, 0); +x_80 = lean_ctor_get(x_73, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_73); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +return x_81; +} +} +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_82 = lean_ctor_get(x_67, 0); +x_83 = lean_ctor_get(x_67, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_67); +x_84 = lean_box(0); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_Elab_Tactic_replaceMainGoal(x_85, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_88 = x_86; +} else { + lean_dec_ref(x_86); + x_88 = lean_box(0); +} +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_88; +} +lean_ctor_set(x_89, 0, x_82); +lean_ctor_set(x_89, 1, x_87); +return x_89; +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_82); +x_90 = lean_ctor_get(x_86, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_86, 1); +lean_inc(x_91); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_92 = x_86; +} else { + lean_dec_ref(x_86); + x_92 = lean_box(0); +} +if (lean_is_scalar(x_92)) { + x_93 = lean_alloc_ctor(1, 2, 0); +} else { + x_93 = x_92; +} +lean_ctor_set(x_93, 0, x_90); +lean_ctor_set(x_93, 1, x_91); +return x_93; +} +} +} +else +{ +uint8_t x_94; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_94 = !lean_is_exclusive(x_66); +if (x_94 == 0) +{ +return x_66; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_66, 0); +x_96 = lean_ctor_get(x_66, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_66); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; +} +} +} +} +else +{ +uint8_t x_98; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +x_98 = !lean_is_exclusive(x_12); +if (x_98 == 0) +{ +return x_12; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_12, 0); +x_100 = lean_ctor_get(x_12, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_12); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; } } } @@ -29651,8 +29872,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep(lean_object* x_1 _start: { lean_object* x_13; lean_object* x_14; -x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed), 10, 1); -lean_closure_set(x_13, 0, x_2); +lean_inc(x_1); +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed), 11, 2); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_2); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -29780,16 +30003,17 @@ return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_12; +x_12 = l_Lean_Elab_Tactic_evalIntro_introStep___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_11; +lean_dec(x_1); +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -30401,7 +30625,8 @@ x_47 = lean_ctor_get(x_44, 1); x_48 = lean_ctor_get(x_46, 0); lean_inc(x_48); lean_dec(x_46); -x_49 = lean_environment_main_module(x_48); +x_49 = l_Lean_Environment_mainModule(x_48); +lean_dec(x_48); x_50 = l_Lean_Elab_Tactic_evalIntro___lambda__2___closed__1; lean_inc(x_42); lean_ctor_set_tag(x_44, 2); @@ -30541,7 +30766,8 @@ lean_dec(x_44); x_112 = lean_ctor_get(x_110, 0); lean_inc(x_112); lean_dec(x_110); -x_113 = lean_environment_main_module(x_112); +x_113 = l_Lean_Environment_mainModule(x_112); +lean_dec(x_112); x_114 = l_Lean_Elab_Tactic_evalIntro___lambda__2___closed__1; lean_inc(x_42); x_115 = lean_alloc_ctor(2, 2, 0); @@ -30698,7 +30924,8 @@ x_184 = lean_ctor_get(x_181, 1); x_185 = lean_ctor_get(x_183, 0); lean_inc(x_185); lean_dec(x_183); -x_186 = lean_environment_main_module(x_185); +x_186 = l_Lean_Environment_mainModule(x_185); +lean_dec(x_185); x_187 = l_Lean_Elab_Tactic_evalIntro___lambda__2___closed__1; lean_inc(x_179); lean_ctor_set_tag(x_181, 2); @@ -30838,7 +31065,8 @@ lean_dec(x_181); x_249 = lean_ctor_get(x_247, 0); lean_inc(x_249); lean_dec(x_247); -x_250 = lean_environment_main_module(x_249); +x_250 = l_Lean_Environment_mainModule(x_249); +lean_dec(x_249); x_251 = l_Lean_Elab_Tactic_evalIntro___lambda__2___closed__1; lean_inc(x_179); x_252 = lean_alloc_ctor(2, 2, 0); @@ -30997,7 +31225,8 @@ x_322 = lean_ctor_get(x_319, 1); x_323 = lean_ctor_get(x_321, 0); lean_inc(x_323); lean_dec(x_321); -x_324 = lean_environment_main_module(x_323); +x_324 = l_Lean_Environment_mainModule(x_323); +lean_dec(x_323); x_325 = l_Lean_Elab_Tactic_evalIntro___lambda__2___closed__1; lean_inc(x_317); lean_ctor_set_tag(x_319, 2); @@ -31137,7 +31366,8 @@ lean_dec(x_319); x_387 = lean_ctor_get(x_385, 0); lean_inc(x_387); lean_dec(x_385); -x_388 = lean_environment_main_module(x_387); +x_388 = l_Lean_Environment_mainModule(x_387); +lean_dec(x_387); x_389 = l_Lean_Elab_Tactic_evalIntro___lambda__2___closed__1; lean_inc(x_317); x_390 = lean_alloc_ctor(2, 2, 0); @@ -48782,7 +49012,8 @@ x_23 = lean_ctor_get(x_20, 1); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); +x_25 = l_Lean_Environment_mainModule(x_24); +lean_dec(x_24); x_26 = l_Lean_Elab_Tactic_evalRunTac___closed__7; x_27 = l_Lean_addMacroScope(x_25, x_26, x_19); x_28 = l_Lean_Elab_Tactic_evalRunTac___closed__6; @@ -48866,7 +49097,8 @@ lean_dec(x_20); x_48 = lean_ctor_get(x_46, 0); lean_inc(x_48); lean_dec(x_46); -x_49 = lean_environment_main_module(x_48); +x_49 = l_Lean_Environment_mainModule(x_48); +lean_dec(x_48); x_50 = l_Lean_Elab_Tactic_evalRunTac___closed__7; x_51 = l_Lean_addMacroScope(x_49, x_50, x_19); x_52 = l_Lean_Elab_Tactic_evalRunTac___closed__6; diff --git a/stage0/stdlib/Lean/Elab/Tactic/Config.c b/stage0/stdlib/Lean/Elab/Tactic/Config.c index a00fde5def87..f830f06138ce 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Config.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Config.c @@ -392,7 +392,7 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Confi static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__244; static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__252; static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__53; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandConfigElab__1___closed__1; static lean_object* l_Lean_Elab_Tactic_configElab___closed__10; static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__141; @@ -3081,7 +3081,8 @@ x_34 = lean_ctor_get(x_31, 1); x_35 = lean_ctor_get(x_33, 0); lean_inc(x_35); lean_dec(x_33); -x_36 = lean_environment_main_module(x_35); +x_36 = l_Lean_Environment_mainModule(x_35); +lean_dec(x_35); x_37 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__1___lambda__3___closed__11; lean_inc(x_28); lean_inc(x_36); @@ -3142,7 +3143,8 @@ lean_dec(x_31); x_61 = lean_ctor_get(x_59, 0); lean_inc(x_61); lean_dec(x_59); -x_62 = lean_environment_main_module(x_61); +x_62 = l_Lean_Environment_mainModule(x_61); +lean_dec(x_61); x_63 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__1___lambda__3___closed__11; lean_inc(x_28); lean_inc(x_62); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Doc.c b/stage0/stdlib/Lean/Elab/Tactic/Doc.c index a895ef8cbe3c..3b62039da64b 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Doc.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Doc.c @@ -77,7 +77,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Doc_0__Lean_Elab_Tactic_Do static lean_object* l_Lean_Elab_Tactic_Doc_elabTacticExtension___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Doc_elabTacticExtension_declRange__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Doc_allTacticDocs___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Doc_allTacticDocs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -235,7 +235,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Elab_Tactic_Do LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Doc_elabTacticExtension___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Elab_Tactic_Doc_allTacticDocs___spec__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Doc_elabRegisterTacticTag___lambda__1___closed__1; -lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Doc_0__Lean_Elab_Tactic_Doc_getFirstTk___closed__6; static lean_object* l___private_Lean_Elab_Tactic_Doc_0__Lean_Elab_Tactic_Doc_getFirstTk___closed__15; @@ -269,6 +268,7 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getDocStringText___at_Lean_Elab_Tactic_Doc_elabRegisterTacticTag___spec__1___closed__2; 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_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Doc_elabTacticExtension_declRange__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Doc_0__Lean_Elab_Tactic_Doc_getFirstTk(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFnArgs(lean_object*); @@ -7392,11 +7392,10 @@ lean_dec(x_7); x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); -x_11 = lean_ctor_get(x_10, 1); +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -x_12 = l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(x_11, x_1); -lean_inc(x_1); -x_13 = lean_environment_find(x_10, x_1); +x_12 = l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(x_11, x_1); +x_13 = l_Lean_Environment_find_x3f(x_10, x_1); if (lean_obj_tag(x_12) == 0) { lean_object* x_114; @@ -10838,9 +10837,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Doc_allTacticDocs___lambda__2(lean_o _start: { lean_object* x_11; -lean_inc(x_2); lean_inc(x_1); -x_11 = lean_environment_find(x_1, x_2); +x_11 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_11) == 0) { uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; diff --git a/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c b/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c index aae1da32eb2b..0192ac40c75a 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c +++ b/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c @@ -243,7 +243,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tact static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalWithReducibleAndInstances_declRange__1___closed__4; uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabDecideConfig(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); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalWithReducibleAndInstances_declRange__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDecideCore_diagnose___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDecideCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -253,7 +252,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyLikeTactic___lambda__1(lean static lean_object* l_Lean_Elab_Tactic_evalDecideCore_diagnose___lambda__4___closed__28; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_elabNativeDecideCoreUnsafe___boxed(lean_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_Parser_Tactic_getConfigItems(lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_runTermElab___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalRename___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_ElabTerm___hyg_6166____closed__1; @@ -329,6 +327,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalWithReducible(lean_object*, lean lean_object* l_Lean_Elab_Term_elabTerm___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_evalDecideCore_diagnose___lambda__4___closed__8; static lean_object* l_Lean_Elab_Tactic_evalRename___closed__2; +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); lean_object* l_Lean_MessageData_andList(lean_object*); @@ -720,6 +719,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_sortMVarIdsByIndex___at_Lean_Elab_Ta lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalDecideCore_diagnose___lambda__5___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalNativeDecide_declRange__1___closed__6; +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Elab_Tactic_evalDecideCore_diagnose___lambda__4___closed__23; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDecideCore_diagnose___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoalUsing___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*); @@ -17724,7 +17724,7 @@ lean_dec(x_12); x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); lean_dec(x_13); -x_16 = l_Lean_Kernel_isDiagnosticsEnabled(x_15); +x_16 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_15); lean_dec(x_15); if (x_16 == 0) { @@ -17774,7 +17774,7 @@ lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean x_22 = lean_ctor_get(x_19, 0); x_23 = lean_ctor_get(x_19, 4); lean_dec(x_23); -x_24 = l_Lean_Kernel_enableDiag(x_22, x_11); +x_24 = l_Lean_Kernel_Environment_enableDiag(x_22, x_11); x_25 = l_Lean_Elab_Tactic_evalDecideCore_diagnose___lambda__3___closed__2; lean_ctor_set(x_19, 4, x_25); lean_ctor_set(x_19, 0, x_24); @@ -17804,7 +17804,7 @@ lean_inc(x_32); lean_inc(x_31); lean_inc(x_30); lean_dec(x_19); -x_37 = l_Lean_Kernel_enableDiag(x_30, x_11); +x_37 = l_Lean_Kernel_Environment_enableDiag(x_30, x_11); x_38 = l_Lean_Elab_Tactic_evalDecideCore_diagnose___lambda__3___closed__2; x_39 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_39, 0, x_37); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Ext.c b/stage0/stdlib/Lean/Elab/Tactic/Ext.c index a422fddfa855..ceb46925279b 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Ext.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Ext.c @@ -175,7 +175,7 @@ lean_object* l_Lean_Elab_Tactic_getMainGoal(lean_object*, lean_object*, lean_obj lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___at_Lean_Elab_Tactic_Ext_extCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__19; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_ReaderT_read___at_Lean_Macro_instMonadRefMacroM___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__38; @@ -611,7 +611,7 @@ static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___cl static lean_object* l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3663____lambda__6___closed__8; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___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*); @@ -6317,8 +6317,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -6359,8 +6358,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -7013,7 +7011,8 @@ x_17 = lean_ctor_get(x_14, 1); x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); -x_19 = lean_environment_main_module(x_18); +x_19 = l_Lean_Environment_mainModule(x_18); +lean_dec(x_18); x_20 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__9; lean_inc(x_12); lean_ctor_set_tag(x_14, 2); @@ -7383,7 +7382,8 @@ lean_dec(x_14); x_190 = lean_ctor_get(x_188, 0); lean_inc(x_190); lean_dec(x_188); -x_191 = lean_environment_main_module(x_190); +x_191 = l_Lean_Environment_mainModule(x_190); +lean_dec(x_190); x_192 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__9; lean_inc(x_12); x_193 = lean_alloc_ctor(2, 2, 0); @@ -12720,7 +12720,8 @@ x_37 = lean_ctor_get(x_34, 1); x_38 = lean_ctor_get(x_36, 3); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_main_module(x_8); +x_39 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_40 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_40, 0, x_22); lean_ctor_set(x_40, 1, x_39); @@ -12901,7 +12902,8 @@ lean_dec(x_34); x_89 = lean_ctor_get(x_87, 3); lean_inc(x_89); lean_dec(x_87); -x_90 = lean_environment_main_module(x_8); +x_90 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_91 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_91, 0, x_22); lean_ctor_set(x_91, 1, x_90); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Grind.c b/stage0/stdlib/Lean/Elab/Tactic/Grind.c index 3fc44a87b94b..aceacd53aa78 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Grind.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Grind.c @@ -13,21 +13,28 @@ #ifdef __cplusplus extern "C" { #endif +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalGrind__1(lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4; static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__5; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_grind___closed__2; +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2(lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2; +static lean_object* l_Lean_Elab_Tactic_evalGrind___closed__2; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__4; lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -35,14 +42,15 @@ lean_object* l_Lean_indentD(lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__2; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__4; lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_mkAuxName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalGrind(lean_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_grind___closed__1; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_instMonadTermElabM; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); @@ -50,10 +58,10 @@ static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elab lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__6; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalGrind___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__3; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_elabGrindPattern___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -67,103 +75,124 @@ lean_object* l_Lean_Exception_toMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_elabGrindConfig___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__1; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(lean_object*); uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); +lean_object* l_Lean_Syntax_getNumArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13(lean_object*); lean_object* l_Lean_Parser_Tactic_getConfigItems(lean_object*); lean_object* l_List_filterMapTR_go___at_Lean_preprocessSyntaxAndResolve___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__9; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalGrind___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_elabGrindConfig___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_abstract(lean_object*, lean_object*); +lean_object* l_Lean_Meta_withLCtx___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_filterFieldList___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_preprocessPattern(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__1; -lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Command_commandElabAttribute; static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__4; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__2; -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__8; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__3; lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__5; +lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5; lean_object* l_Lean_Elab_goalsToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__4; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; lean_object* l_Lean_Elab_Tactic_withMainContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6; lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabGrindPattern___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_getDeclName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3(lean_object*); lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigItemViews(lean_object*); static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__3; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__5; +static lean_object* l_Lean_Elab_Tactic_evalGrind___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__6; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__2; +static lean_object* l_Lean_Elab_Tactic_evalGrind___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__6; +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__3; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15; lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_resolveLocalName_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__1; extern lean_object* l_Std_Format_defWidth; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14; static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__3; static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux___rarg(uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___closed__4; -lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16; +lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Tactic_elabGrindPattern___spec__3___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1(lean_object*); static lean_object* l_panic___at_Lean_Elab_Tactic_elabGrindPattern___spec__10___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__4; +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_Elab_Tactic_elabGrindPattern___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7___closed__1; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__1; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1; lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalGrind___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindPattern___closed__4; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_elabTermAndSynthesize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +static lean_object* l_Lean_Elab_Tactic_evalGrind___closed__5; LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Tactic_elabGrindPattern___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___boxed(lean_object*, lean_object*); lean_object* l_List_toString___at_Lean_ensureNoOverload___spec__2(lean_object*); static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__3; @@ -182,29 +211,39 @@ size_t lean_array_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__6; lean_object* l_Lean_Elab_Term_addTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__2___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Tactic_elabGrindPattern___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___rarg(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10; lean_object* l_Lean_Elab_Tactic_liftMetaFinishingTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_Elab_Tactic_elabGrindPattern___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8; lean_object* l_Lean_Meta_evalExpr_x27___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindPattern(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__1; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__2; uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addAndCompile(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expandDeclId___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__5; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabGrindPattern___spec__6___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__3; lean_object* l_Lean_MessageData_ofName(lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalApplyRfl___closed__7; +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3; lean_object* l_List_filterTR_loop___at_Lean_filterFieldList___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18; static lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__1; static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -675,18 +714,17 @@ return x_18; static lean_object* _init_l_Lean_Elab_Tactic_elabGrindConfig___lambda__3___closed__1() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = 0; -x_2 = lean_unsigned_to_nat(100u); -x_3 = lean_unsigned_to_nat(5u); -x_4 = lean_unsigned_to_nat(1000u); -x_5 = lean_alloc_ctor(0, 4, 1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_3); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_uint8(x_5, sizeof(void*)*4, x_1); -return x_5; +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = lean_unsigned_to_nat(5u); +x_2 = lean_unsigned_to_nat(1000u); +x_3 = 1; +x_4 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_1); +lean_ctor_set(x_4, 2, x_1); +lean_ctor_set(x_4, 3, x_2); +lean_ctor_set_uint8(x_4, sizeof(void*)*4, x_3); +return x_4; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabGrindConfig___lambda__3(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { @@ -1746,22 +1784,6 @@ return x_17; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); -return x_1; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); -return x_1; -} -} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -1799,25 +1821,36 @@ lean_inc(x_1); x_19 = l_Lean_Elab_Term_elabTerm(x_15, x_1, x_18, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_19) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_20; lean_object* x_21; lean_object* x_22; x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_22 = l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(x_6, x_7, x_8, x_9, x_10, x_11, x_21); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_22, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); lean_dec(x_22); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1; -x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2; +x_24 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_27 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_25, x_26, x_8, x_9, x_10, x_11, x_24); +x_27 = l_Lean_Meta_Grind_preprocessPattern(x_25, x_18, x_8, x_9, x_10, x_11, x_26); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; @@ -1870,6 +1903,7 @@ return x_38; else { uint8_t x_39; +lean_dec(x_20); lean_dec(x_17); lean_dec(x_11); lean_dec(x_10); @@ -1878,19 +1912,19 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_39 = !lean_is_exclusive(x_19); +x_39 = !lean_is_exclusive(x_22); if (x_39 == 0) { -return x_19; +return x_22; } else { lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_19, 0); -x_41 = lean_ctor_get(x_19, 1); +x_40 = lean_ctor_get(x_22, 0); +x_41 = lean_ctor_get(x_22, 1); lean_inc(x_41); lean_inc(x_40); -lean_dec(x_19); +lean_dec(x_22); x_42 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -1898,6 +1932,37 @@ return x_42; } } } +else +{ +uint8_t x_43; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_19); +if (x_43 == 0) +{ +return x_19; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_19, 0); +x_45 = lean_ctor_get(x_19, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_19); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} } } LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Elab_Tactic_elabGrindPattern___spec__13___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { @@ -2535,277 +2600,942 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_grind(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_9; +lean_object* x_10; +lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_9 = l_Lean_Meta_Grind_main(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_10 = l_Lean_Meta_Grind_main(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = l_List_isEmpty___rarg(x_11); -if (x_13 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = l_List_isEmpty___rarg(x_12); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_free_object(x_9); -x_14 = l_Lean_Elab_goalsToMessageData(x_11); -x_15 = l_Lean_Elab_Tactic_grind___closed__2; -x_16 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -x_17 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; -x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_18, x_4, x_5, x_6, x_7, x_12); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_free_object(x_10); +x_15 = l_Lean_Elab_goalsToMessageData(x_12); +x_16 = l_Lean_Elab_Tactic_grind___closed__2; +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_19, x_5, x_6, x_7, x_8, x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_19; +return x_20; } else { -lean_object* x_20; -lean_dec(x_11); +lean_object* x_21; +lean_dec(x_12); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_20 = lean_box(0); -lean_ctor_set(x_9, 0, x_20); -return x_9; +x_21 = lean_box(0); +lean_ctor_set(x_10, 0, x_21); +return x_10; } } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_9, 0); -x_22 = lean_ctor_get(x_9, 1); +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_10, 0); +x_23 = lean_ctor_get(x_10, 1); +lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_9); -x_23 = l_List_isEmpty___rarg(x_21); -if (x_23 == 0) +lean_dec(x_10); +x_24 = l_List_isEmpty___rarg(x_22); +if (x_24 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_24 = l_Lean_Elab_goalsToMessageData(x_21); -x_25 = l_Lean_Elab_Tactic_grind___closed__2; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -x_27 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_28, x_4, x_5, x_6, x_7, x_22); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_25 = l_Lean_Elab_goalsToMessageData(x_22); +x_26 = l_Lean_Elab_Tactic_grind___closed__2; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_Elab_Tactic_elabGrindConfig___lambda__1___closed__6; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_29, x_5, x_6, x_7, x_8, x_23); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_29; +return x_30; } else { -lean_object* x_30; lean_object* x_31; -lean_dec(x_21); +lean_object* x_31; lean_object* x_32; +lean_dec(x_22); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_30 = lean_box(0); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_22); -return x_31; +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_23); +return x_32; } } } else { -uint8_t x_32; +uint8_t x_33; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_32 = !lean_is_exclusive(x_9); -if (x_32 == 0) +x_33 = !lean_is_exclusive(x_10); +if (x_33 == 0) { -return x_9; +return x_10; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_9, 0); -x_34 = lean_ctor_get(x_9, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_10, 0); +x_35 = lean_ctor_get(x_10, 1); +lean_inc(x_35); lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_9); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_dec(x_10); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_elabGrindPattern___spec__1___rarg___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +lean_inc(x_10); +x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_9; -x_9 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg), 1, 0); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_14, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_9; -x_9 = l_Lean_Elab_Tactic_grind(x_3, x_1, x_2, x_4, x_5, x_6, x_7, x_8); -return x_9; -} +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_12, 1, x_18); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 0, x_12); +return x_16; } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_12, 1, x_19); +lean_ctor_set(x_12, 0, x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; -x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; -x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; -x_4 = l_Lean_Elab_Tactic_evalApplyRfl___closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_12); +x_24 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_22, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); } +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_11); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(1, 2, 0); +} else { + x_29 = x_27; + lean_ctor_set_tag(x_29, 1); +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("optConfig", 9, 9); -return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__4() { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; -x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; -x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; -x_4 = l_Lean_Elab_Tactic_evalApplyRfl___closed__3; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg___boxed), 8, 0); +return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__5() { +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("The `grind` tactic is experimental and still under development. Avoid using it in production projects", 101, 101); -return x_1; -} +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_1, 0); +x_10 = l_Lean_stringToMessageData(x_9); +x_11 = l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalApplyRfl___closed__5; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_8); +return x_13; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__7() { +} +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalApplyRfl___closed__6; -x_2 = l_Lean_MessageData_ofFormat(x_1); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg___boxed), 8, 0); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__8() { +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_6, 2); +x_14 = lean_eval_const(x_12, x_13, x_1); +lean_dec(x_12); +x_15 = l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg___boxed), 8, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__3___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_ofExcept___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__2___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___rarg), 1, 0); +return x_9; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("_grind", 6, 6); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___boxed), 8, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("GoalM", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2; +x_3 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__2; +x_4 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Unit", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5; +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8; +x_3 = l_Lean_Expr_app___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__9() { -_start: +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13; +x_3 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11; +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_evalConst___at___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback_unsafe__1___spec__1___rarg___boxed), 8, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_grind_fallback", 15, 15); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_9 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_box(0); +x_14 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17; +x_15 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermAndSynthesize), 9, 2); +lean_closure_set(x_15, 0, x_12); +lean_closure_set(x_15, 1, x_14); +x_16 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15; +x_17 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_18 = l_Lean_Meta_withLCtx___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__3___rarg(x_16, x_17, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18; +if (lean_obj_tag(x_19) == 4) +{ +lean_object* x_22; lean_object* x_23; +lean_free_object(x_1); +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_apply_8(x_21, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20; +lean_inc(x_2); +x_25 = l_Lean_Elab_Term_mkAuxName(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9; +lean_inc(x_26); +x_29 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_13); +lean_ctor_set(x_29, 2, x_28); +lean_inc(x_26); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_26); +lean_ctor_set(x_30, 1, x_13); +x_31 = lean_box(0); +x_32 = 1; +x_33 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_33, 0, x_29); +lean_ctor_set(x_33, 1, x_19); +lean_ctor_set(x_33, 2, x_31); +lean_ctor_set(x_33, 3, x_30); +lean_ctor_set_uint8(x_33, sizeof(void*)*4, x_32); +lean_ctor_set(x_1, 0, x_33); +lean_inc(x_7); +lean_inc(x_6); +x_34 = l_Lean_addAndCompile(x_1, x_6, x_7, x_27); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_apply_8(x_21, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_35); +return x_36; +} +else +{ +uint8_t x_37; +lean_dec(x_26); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_34); +if (x_37 == 0) +{ +return x_34; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_34, 0); +x_39 = lean_ctor_get(x_34, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_34); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +else +{ +uint8_t x_41; +lean_dec(x_19); +lean_free_object(x_1); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_41 = !lean_is_exclusive(x_25); +if (x_41 == 0) +{ +return x_25; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_25, 0); +x_43 = lean_ctor_get(x_25, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_25); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +} +else +{ +uint8_t x_45; +lean_free_object(x_1); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_45 = !lean_is_exclusive(x_18); +if (x_45 == 0) +{ +return x_18; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_18, 0); +x_47 = lean_ctor_get(x_18, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_18); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalApplyRfl___closed__8; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_49 = lean_ctor_get(x_1, 0); +lean_inc(x_49); +lean_dec(x_1); +x_50 = lean_box(0); +x_51 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17; +x_52 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermAndSynthesize), 9, 2); +lean_closure_set(x_52, 0, x_49); +lean_closure_set(x_52, 1, x_51); +x_53 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15; +x_54 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_55 = l_Lean_Meta_withLCtx___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__3___rarg(x_53, x_54, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18; +if (lean_obj_tag(x_56) == 4) +{ +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 0); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_apply_8(x_58, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_57); +return x_60; } +else +{ +lean_object* x_61; lean_object* x_62; +x_61 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20; +lean_inc(x_2); +x_62 = l_Lean_Elab_Term_mkAuxName(x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_57); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9; +lean_inc(x_63); +x_66 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_66, 0, x_63); +lean_ctor_set(x_66, 1, x_50); +lean_ctor_set(x_66, 2, x_65); +lean_inc(x_63); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_63); +lean_ctor_set(x_67, 1, x_50); +x_68 = lean_box(0); +x_69 = 1; +x_70 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_70, 0, x_66); +lean_ctor_set(x_70, 1, x_56); +lean_ctor_set(x_70, 2, x_68); +lean_ctor_set(x_70, 3, x_67); +lean_ctor_set_uint8(x_70, sizeof(void*)*4, x_69); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +lean_inc(x_7); +lean_inc(x_6); +x_72 = l_Lean_addAndCompile(x_71, x_6, x_7, x_64); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_72, 1); +lean_inc(x_73); +lean_dec(x_72); +x_74 = lean_apply_8(x_58, x_63, x_2, x_3, x_4, x_5, x_6, x_7, x_73); +return x_74; } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalApplyRfl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_Elab_Tactic_evalApplyRfl___closed__2; -lean_inc(x_1); -x_12 = l_Lean_Syntax_isOfKind(x_1, x_11); -if (x_12 == 0) +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_75 = lean_ctor_get(x_72, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_72, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_77 = x_72; +} else { + lean_dec_ref(x_72); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +else { -lean_object* x_13; -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_56); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(x_10); -return x_13; +x_79 = lean_ctor_get(x_62, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_62, 1); +lean_inc(x_80); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_81 = x_62; +} else { + lean_dec_ref(x_62); + x_81 = lean_box(0); +} +if (lean_is_scalar(x_81)) { + x_82 = lean_alloc_ctor(1, 2, 0); +} else { + x_82 = x_81; +} +lean_ctor_set(x_82, 0, x_79); +lean_ctor_set(x_82, 1, x_80); +return x_82; +} +} } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(1u); -x_15 = l_Lean_Syntax_getArg(x_1, x_14); -x_16 = l_Lean_Elab_Tactic_evalApplyRfl___closed__4; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_83 = lean_ctor_get(x_55, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_55, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_85 = x_55; +} else { + lean_dec_ref(x_55); + x_85 = lean_box(0); +} +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); +} else { + x_86 = x_85; +} +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_9); +lean_object* x_9; +x_9 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___elambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2814,55 +3544,153 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___rarg(x_10); -return x_18; +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalGrind___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Elab_Tactic_grind(x_4, x_1, x_2, x_3, x_5, x_6, x_7, x_8, x_9); +return x_10; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalGrind___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("The `grind` tactic is experimental and still under development. Avoid using it in production projects", 101, 101); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalGrind___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_evalGrind___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalGrind___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_evalGrind___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalGrind___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_grind", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalGrind___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_evalGrind___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalGrind(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_11 = lean_unsigned_to_nat(1u); +x_12 = l_Lean_Syntax_getArg(x_1, x_11); +x_13 = lean_unsigned_to_nat(2u); +x_14 = l_Lean_Syntax_getArg(x_1, x_13); +x_15 = l_Lean_Syntax_getNumArgs(x_14); +x_16 = lean_nat_dec_eq(x_15, x_13); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_54; +lean_dec(x_14); +x_54 = lean_box(0); +x_17 = x_54; +goto block_53; } else { -lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_19 = l_Lean_Elab_Tactic_evalApplyRfl___closed__7; -x_20 = 1; +lean_object* x_55; lean_object* x_56; +x_55 = l_Lean_Syntax_getArg(x_14, x_11); +lean_dec(x_14); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_17 = x_56; +goto block_53; +} +block_53: +{ +lean_object* x_18; +lean_inc(x_9); lean_inc(x_8); -x_21 = l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2(x_1, x_19, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Elab_Term_getDeclName_x3f(x_4, x_5, x_6, x_7, x_8, x_9, x_22); -x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_18 = l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback(x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Tactic_evalGrind___closed__3; +x_22 = 1; +lean_inc(x_8); +x_23 = l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2(x_1, x_21, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); lean_dec(x_23); +x_25 = l_Lean_Elab_Term_getDeclName_x3f(x_4, x_5, x_6, x_7, x_8, x_9, x_24); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_26 = l_Lean_Elab_Tactic_elabGrindConfig(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -if (lean_obj_tag(x_24) == 0) -{ +x_28 = l_Lean_Elab_Tactic_elabGrindConfig(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_evalApplyRfl___closed__9; -x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 8, 2); -lean_closure_set(x_30, 0, x_27); -lean_closure_set(x_30, 1, x_29); -x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); -lean_closure_set(x_31, 0, x_30); -x_32 = l_Lean_Elab_Tactic_withMainContext___rarg(x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); -return x_32; +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Elab_Tactic_evalGrind___closed__5; +x_32 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalGrind___lambda__1), 9, 3); +lean_closure_set(x_32, 0, x_29); +lean_closure_set(x_32, 1, x_31); +lean_closure_set(x_32, 2, x_19); +x_33 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); +lean_closure_set(x_33, 0, x_32); +x_34 = l_Lean_Elab_Tactic_withMainContext___rarg(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_30); +return x_34; } else { -uint8_t x_33; +uint8_t x_35; +lean_dec(x_19); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2871,51 +3699,53 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_33 = !lean_is_exclusive(x_26); -if (x_33 == 0) +x_35 = !lean_is_exclusive(x_28); +if (x_35 == 0) { -return x_26; +return x_28; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_26, 0); -x_35 = lean_ctor_get(x_26, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_26); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_28, 0); +x_37 = lean_ctor_get(x_28, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_28); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } } else { -if (lean_obj_tag(x_26) == 0) +if (lean_obj_tag(x_28) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_37 = lean_ctor_get(x_24, 0); -lean_inc(x_37); -lean_dec(x_24); -x_38 = lean_ctor_get(x_26, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_26, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_39 = lean_ctor_get(x_26, 0); lean_inc(x_39); lean_dec(x_26); -x_40 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl___lambda__1), 8, 2); -lean_closure_set(x_40, 0, x_38); -lean_closure_set(x_40, 1, x_37); -x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); -lean_closure_set(x_41, 0, x_40); -x_42 = l_Lean_Elab_Tactic_withMainContext___rarg(x_41, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_39); -return x_42; +x_40 = lean_ctor_get(x_28, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_28, 1); +lean_inc(x_41); +lean_dec(x_28); +x_42 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalGrind___lambda__1), 9, 3); +lean_closure_set(x_42, 0, x_40); +lean_closure_set(x_42, 1, x_39); +lean_closure_set(x_42, 2, x_19); +x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_liftMetaFinishingTactic), 10, 1); +lean_closure_set(x_43, 0, x_42); +x_44 = l_Lean_Elab_Tactic_withMainContext___rarg(x_43, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_41); +return x_44; } else { -uint8_t x_43; -lean_dec(x_24); +uint8_t x_45; +lean_dec(x_26); +lean_dec(x_19); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2924,35 +3754,32 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_43 = !lean_is_exclusive(x_26); -if (x_43 == 0) +x_45 = !lean_is_exclusive(x_28); +if (x_45 == 0) { -return x_26; +return x_28; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_26, 0); -x_45 = lean_ctor_get(x_26, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_26); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} -} -} +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_28, 0); +x_47 = lean_ctor_get(x_28, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_28); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalApplyRfl___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +uint8_t x_49; +lean_dec(x_12); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2960,31 +3787,78 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +x_49 = !lean_is_exclusive(x_18); +if (x_49 == 0) +{ +return x_18; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_18, 0); +x_51 = lean_ctor_get(x_18, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_18); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalGrind___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Elab_Tactic_evalGrind(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_1); -return x_9; +return x_11; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; +x_2 = l_Lean_Elab_Tactic_elabGrindPattern___closed__1; +x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Tactic_evalGrind__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_Tactic_evalApplyRfl__1___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("evalApplyRfl", 12, 12); +x_1 = lean_mk_string_unchecked("evalGrind", 9, 9); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Grind___hyg_5____closed__1; x_2 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__1; x_3 = l___regBuiltin_Lean_Elab_Tactic_elabGrindPattern__1___closed__2; -x_4 = l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1; +x_4 = l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__5() { _start: { lean_object* x_1; @@ -2992,22 +3866,22 @@ x_1 = l_Lean_Elab_Tactic_tacticElabAttribute; return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalApplyRfl), 10, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalGrind___boxed), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalGrind__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__3; -x_3 = l_Lean_Elab_Tactic_evalApplyRfl___closed__2; -x_4 = l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2; -x_5 = l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__5; +x_3 = l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__4; +x_5 = l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__6; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } @@ -3109,10 +3983,6 @@ l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___clo lean_mark_persistent(l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__5); l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6 = _init_l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6(); lean_mark_persistent(l_Lean_ensureNonAmbiguous___at_Lean_Elab_Tactic_elabGrindPattern___spec__9___closed__6); -l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabGrindPattern___spec__12___closed__2); l_Lean_Elab_Tactic_elabGrindPattern___closed__1 = _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_elabGrindPattern___closed__1); l_Lean_Elab_Tactic_elabGrindPattern___closed__2 = _init_l_Lean_Elab_Tactic_elabGrindPattern___closed__2(); @@ -3144,33 +4014,69 @@ lean_dec_ref(res); lean_mark_persistent(l_Lean_Elab_Tactic_grind___closed__1); l_Lean_Elab_Tactic_grind___closed__2 = _init_l_Lean_Elab_Tactic_grind___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_grind___closed__2); -l_Lean_Elab_Tactic_evalApplyRfl___closed__1 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__1); -l_Lean_Elab_Tactic_evalApplyRfl___closed__2 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__2); -l_Lean_Elab_Tactic_evalApplyRfl___closed__3 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__3); -l_Lean_Elab_Tactic_evalApplyRfl___closed__4 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__4); -l_Lean_Elab_Tactic_evalApplyRfl___closed__5 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__5); -l_Lean_Elab_Tactic_evalApplyRfl___closed__6 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__6); -l_Lean_Elab_Tactic_evalApplyRfl___closed__7 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__7); -l_Lean_Elab_Tactic_evalApplyRfl___closed__8 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__8); -l_Lean_Elab_Tactic_evalApplyRfl___closed__9 = _init_l_Lean_Elab_Tactic_evalApplyRfl___closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalApplyRfl___closed__9); -l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__1); -l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__2); -l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__3 = _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__3); -l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__4 = _init_l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1___closed__4); -if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_evalApplyRfl__1(lean_io_mk_world()); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__1); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__2); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__3); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__4); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__5); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__6); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__7); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__8); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__9); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__10); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__11); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__12); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__13); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__14); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__15); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__16); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__17); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__18); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__19); +l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20 = _init_l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Grind_0__Lean_Elab_Tactic_elabFallback___closed__20); +l_Lean_Elab_Tactic_evalGrind___closed__1 = _init_l_Lean_Elab_Tactic_evalGrind___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalGrind___closed__1); +l_Lean_Elab_Tactic_evalGrind___closed__2 = _init_l_Lean_Elab_Tactic_evalGrind___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalGrind___closed__2); +l_Lean_Elab_Tactic_evalGrind___closed__3 = _init_l_Lean_Elab_Tactic_evalGrind___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalGrind___closed__3); +l_Lean_Elab_Tactic_evalGrind___closed__4 = _init_l_Lean_Elab_Tactic_evalGrind___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalGrind___closed__4); +l_Lean_Elab_Tactic_evalGrind___closed__5 = _init_l_Lean_Elab_Tactic_evalGrind___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalGrind___closed__5); +l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__1); +l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__2); +l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__3 = _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__3); +l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__4 = _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__4); +l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__5 = _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__5); +l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__6 = _init_l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__6(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalGrind__1___closed__6); +if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_evalGrind__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Induction.c b/stage0/stdlib/Lean/Elab/Tactic/Induction.c index 8051c7745a75..109dc9b34227 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Induction.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Induction.c @@ -13,7 +13,7 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingName_x21(lean_object*); @@ -85,9 +85,11 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_eva LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandInductionAlts_x3f(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__5___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___lambda__2___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14; LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_isMultiAlt(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_promise_new(lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); @@ -96,7 +98,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6___closed__2; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9; lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange__1___closed__2; @@ -107,7 +108,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Elim LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_elabCasesTargets___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Meta_mkGeneralizationForbiddenSet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,7 +124,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_eval static lean_object* l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltVars___boxed(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_evalInduction___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange__1___closed__6; @@ -162,7 +161,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAlt___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_Elab_Tactic_evalInduction___lambda__5(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts(lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__1___closed__1; lean_object* l_Lean_MVarId_setKind(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -171,9 +170,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__5___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__3(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__5___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getNumExplicitFields___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -228,12 +225,10 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_a LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_isHoleRHS(lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__11; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4; uint8_t l_Lean_Expr_isType(lean_object*); lean_object* l_Lean_Syntax_getNumArgs(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_elabCasesTargets___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -265,9 +260,9 @@ lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27(lean_object*); lean_object* l_Lean_Elab_Tactic_getGoals___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_getUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_setMotiveArg___closed__2; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_elabTerm(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_instInhabitedAlt___closed__1; @@ -287,6 +282,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor(uint8_t, le static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__4___closed__2; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___closed__1; static lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__2___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -310,6 +306,7 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_eva LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5; lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_evalInduction___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__4; @@ -325,7 +322,7 @@ lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_elabCasesTargets___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_isHoleRHS___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_elabCasesTargets___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction__1___closed__4; @@ -346,6 +343,7 @@ lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltVars(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_isWildcard___boxed(lean_object*); lean_object* l_Lean_Meta_getCustomEliminator_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -372,14 +370,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___at_Lean_E lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getNumExplicitFields___lambda__1___closed__1; lean_object* l_Lean_Meta_abstractMVars(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_withMainContext___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -396,7 +392,8 @@ extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1(lean_object*); lean_object* l_Lean_MVarId_revert(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763_(lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__9; extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAlt___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -407,6 +404,7 @@ lean_object* l_Lean_Expr_constName_x3f(lean_object*); lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___closed__1; static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__5; +static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -434,6 +432,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabCasesTargets___lambda__2___boxed LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltsOfInductionAlts___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__4(lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCases___lambda__4___boxed(lean_object**); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -445,6 +444,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tac LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__7(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -453,10 +453,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__6___boxed(le static lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_addConstInfo___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_evalTactic_eval___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12; extern lean_object* l_Lean_Language_Snapshot_Diagnostics_empty; static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__7; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__7___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction_checkTargets___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -481,7 +479,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___closed__1; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__8; lean_object* l_Array_append___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -535,12 +532,10 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tact LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltDArrow___boxed(lean_object*); lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__6; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getFType___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___closed__1; @@ -557,14 +552,18 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction__2(lean_o lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltDArrow(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getOptPreTacOfOptInductionAlts(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10; LEAN_EXPORT uint8_t l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__13(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7; lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Generalize_0__Lean_Meta_generalizeCore___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___boxed__const__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -573,22 +572,20 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Elim lean_object* l_Lean_indentExpr(lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); static lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__4___closed__1; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7; lean_object* l_Lean_Meta_getMVarsNoDelayed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6; extern lean_object* l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot; lean_object* l_Lean_Meta_isExprDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getNumExplicitFields___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___boxed(lean_object**); lean_object* l_Lean_Meta_getFVarSetToGeneralize(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ElimApp_setMotiveArg___closed__6; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13; lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__13; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -624,16 +621,16 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tac lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_elabCasesTargets___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1___boxed(lean_object**); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__4; lean_object* l_List_iotaTR(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange__1___closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1; lean_object* l_Array_ofSubarray___rarg(lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandInductionAlts_x3f___spec__1(lean_object*, size_t, size_t); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCases_declRange__1___closed__5; @@ -641,6 +638,7 @@ static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_eva LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Tactic_elabCasesTargets___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getBindingName___boxed(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__3___closed__2; @@ -648,7 +646,6 @@ static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_g LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_elabCasesTargets___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__5___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getFType___boxed(lean_object*); static lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__5; @@ -662,9 +659,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalInduction___lambda__3___boxed(le static lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___closed__4; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8___closed__7; static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6___closed__4; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1; lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_evalGeneralize___spec__9(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___boxed(lean_object**); lean_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); @@ -721,6 +720,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tact static lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__10; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__14; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704_(lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); uint8_t l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2; @@ -736,8 +736,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCases(lean_object*, lean_object* lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___spec__1___boxed(lean_object**); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2; uint8_t l_Lean_Exception_isRuntime(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_ElimApp_mkElimApp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expandDeclId___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -745,6 +743,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Elim LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__8; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___closed__1; @@ -759,7 +758,6 @@ uint64_t l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction_declRange__1___closed__5; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___closed__2; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___lambda__1___closed__2; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* lean_dbg_trace(lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -776,15 +774,18 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_evalCases_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames___spec__4___lambda__1___closed__5; LEAN_EXPORT uint8_t l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___lambda__1(uint8_t, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getUserGeneralizingFVarIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCases_declRange__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAlt___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_evalInduction_checkTargets___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getFType___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3___boxed(lean_object**); @@ -13864,7 +13865,7 @@ lean_inc(x_33); x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); lean_dec(x_32); -x_35 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_30, x_26, x_28, x_10, x_24, x_4, x_25, x_6, x_7, x_2, x_3, x_8, x_9, x_5, x_33, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_34); +x_35 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_30, x_26, x_28, x_10, x_24, x_3, x_25, x_5, x_6, x_2, x_8, x_7, x_9, x_4, x_33, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_34); return x_35; } else @@ -13885,8 +13886,8 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_36 = !lean_is_exclusive(x_32); @@ -13913,7 +13914,7 @@ else { lean_object* x_40; lean_inc(x_28); -x_40 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_30, x_26, x_28, x_10, x_24, x_4, x_25, x_6, x_7, x_2, x_3, x_8, x_9, x_5, x_28, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_29); +x_40 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_30, x_26, x_28, x_10, x_24, x_3, x_25, x_5, x_6, x_2, x_8, x_7, x_9, x_4, x_28, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_29); return x_40; } } @@ -13933,8 +13934,8 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_41 = !lean_is_exclusive(x_27); @@ -14042,7 +14043,7 @@ lean_inc(x_70); x_71 = lean_ctor_get(x_69, 1); lean_inc(x_71); lean_dec(x_69); -x_72 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_67, x_63, x_65, x_10, x_61, x_4, x_62, x_6, x_7, x_2, x_3, x_8, x_9, x_5, x_70, x_12, x_13, x_14, x_15, x_16, x_17, x_60, x_19, x_71); +x_72 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_67, x_63, x_65, x_10, x_61, x_3, x_62, x_5, x_6, x_2, x_8, x_7, x_9, x_4, x_70, x_12, x_13, x_14, x_15, x_16, x_17, x_60, x_19, x_71); return x_72; } else @@ -14063,8 +14064,8 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_73 = lean_ctor_get(x_69, 0); @@ -14093,7 +14094,7 @@ else { lean_object* x_77; lean_inc(x_65); -x_77 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_67, x_63, x_65, x_10, x_61, x_4, x_62, x_6, x_7, x_2, x_3, x_8, x_9, x_5, x_65, x_12, x_13, x_14, x_15, x_16, x_17, x_60, x_19, x_66); +x_77 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__8(x_67, x_63, x_65, x_10, x_61, x_3, x_62, x_5, x_6, x_2, x_8, x_7, x_9, x_4, x_65, x_12, x_13, x_14, x_15, x_16, x_17, x_60, x_19, x_66); return x_77; } } @@ -14113,8 +14114,8 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_78 = lean_ctor_get(x_64, 0); @@ -14493,8 +14494,8 @@ lean_object* x_20 = _args[19]; { lean_object* x_21; x_21 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_21; @@ -14625,8 +14626,8 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_11); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_27 = lean_alloc_ctor(0, 2, 0); @@ -14649,8 +14650,8 @@ lean_inc(x_18); lean_inc(x_17); lean_inc(x_11); lean_inc(x_9); -lean_inc(x_7); -lean_inc(x_5); +lean_inc(x_8); +lean_inc(x_6); lean_inc(x_4); lean_inc(x_3); x_29 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_28, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25); @@ -14681,8 +14682,8 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_11); lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); x_35 = !lean_is_exclusive(x_29); @@ -14743,15 +14744,16 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24) { _start: { -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_10, 1); -x_25 = lean_nat_dec_lt(x_12, x_24); -if (x_25 == 0) +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_11, 1); +x_26 = lean_nat_dec_lt(x_13, x_25); +if (x_26 == 0) { -lean_object* x_26; +lean_object* x_27; +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -14759,56 +14761,56 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_12); +lean_dec(x_13); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_11); -lean_ctor_set(x_26, 1, x_23); -return x_26; +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_12); +lean_ctor_set(x_27, 1, x_24); +return x_27; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_34 = lean_array_fget(x_3, x_12); -x_35 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName(x_34); -lean_inc(x_35); -x_48 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___lambda__1___boxed), 2, 1); -lean_closure_set(x_48, 0, x_35); -x_49 = lean_unsigned_to_nat(0u); -x_50 = l_Array_findFinIdx_x3f_loop___rarg(x_48, x_11, x_49); -if (lean_obj_tag(x_50) == 0) -{ -uint8_t x_51; -x_51 = l_Array_isEmpty___rarg(x_11); -if (x_51 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_35 = lean_array_fget(x_9, x_13); +x_36 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName(x_35); +lean_inc(x_36); +x_49 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___lambda__1___boxed), 2, 1); +lean_closure_set(x_49, 0, x_36); +x_50 = lean_unsigned_to_nat(0u); +x_51 = l_Array_findFinIdx_x3f_loop___rarg(x_49, x_12, x_50); +if (lean_obj_tag(x_51) == 0) { uint8_t x_52; -x_52 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_isWildcard(x_34); +x_52 = l_Array_isEmpty___rarg(x_12); if (x_52 == 0) { -lean_object* x_53; +uint8_t x_53; +x_53 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_isWildcard(x_35); +if (x_53 == 0) +{ +lean_object* x_54; +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_53 = lean_box(0); -x_36 = x_53; -goto block_47; +x_54 = lean_box(0); +x_37 = x_54; +goto block_48; } else { -lean_object* x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_35); -x_54 = lean_box(0); -x_55 = lean_array_size(x_11); -x_56 = 0; -x_57 = lean_box(0); +lean_object* x_55; size_t x_56; size_t x_57; lean_object* x_58; lean_object* x_59; +lean_dec(x_36); +x_55 = lean_box(0); +x_56 = lean_array_size(x_12); +x_57 = 0; +x_58 = lean_box(0); +lean_inc(x_23); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); @@ -14816,28 +14818,28 @@ lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_12); +lean_inc(x_13); +lean_inc(x_9); lean_inc(x_7); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_58 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12, x_11, x_34, x_54, x_11, x_55, x_56, x_57, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -lean_dec(x_11); -if (lean_obj_tag(x_58) == 0) +x_59 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__3(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13, x_12, x_35, x_55, x_12, x_56, x_57, x_58, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +lean_dec(x_12); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___closed__3; -x_27 = x_60; -x_28 = x_59; -goto block_33; +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___closed__3; +x_28 = x_61; +x_29 = x_60; +goto block_34; } else { -uint8_t x_61; +uint8_t x_62; +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -14845,57 +14847,57 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_12); +lean_dec(x_13); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_61 = !lean_is_exclusive(x_58); -if (x_61 == 0) +x_62 = !lean_is_exclusive(x_59); +if (x_62 == 0) { -return x_58; +return x_59; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_58, 0); -x_63 = lean_ctor_get(x_58, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_59, 0); +x_64 = lean_ctor_get(x_59, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_58); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_dec(x_59); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } } else { -lean_object* x_65; +lean_object* x_66; +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_65 = lean_box(0); -x_36 = x_65; -goto block_47; +x_66 = lean_box(0); +x_37 = x_66; +goto block_48; } } else { -uint8_t x_66; -lean_dec(x_35); -x_66 = !lean_is_exclusive(x_50); -if (x_66 == 0) +uint8_t x_67; +lean_dec(x_36); +x_67 = !lean_is_exclusive(x_51); +if (x_67 == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_50, 0); -x_68 = lean_array_fget(x_11, x_67); +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_51, 0); +x_69 = lean_array_fget(x_12, x_68); +lean_inc(x_23); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); @@ -14903,30 +14905,30 @@ lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_12); +lean_inc(x_13); +lean_inc(x_9); lean_inc(x_7); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_69 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12, x_34, x_68, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -if (lean_obj_tag(x_69) == 0) +x_70 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13, x_35, x_69, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +if (lean_obj_tag(x_70) == 0) { -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_69, 1); -lean_inc(x_70); -lean_dec(x_69); -x_71 = l_Array_eraseIdx___rarg(x_11, x_67, lean_box(0)); -lean_ctor_set(x_50, 0, x_71); -x_27 = x_50; -x_28 = x_70; -goto block_33; +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec(x_70); +x_72 = l_Array_eraseIdx___rarg(x_12, x_68, lean_box(0)); +lean_ctor_set(x_51, 0, x_72); +x_28 = x_51; +x_29 = x_71; +goto block_34; } else { -uint8_t x_72; -lean_free_object(x_50); -lean_dec(x_67); +uint8_t x_73; +lean_free_object(x_51); +lean_dec(x_68); +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -14934,40 +14936,40 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_72 = !lean_is_exclusive(x_69); -if (x_72 == 0) +x_73 = !lean_is_exclusive(x_70); +if (x_73 == 0) { -return x_69; +return x_70; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_69, 0); -x_74 = lean_ctor_get(x_69, 1); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_70, 0); +x_75 = lean_ctor_get(x_70, 1); +lean_inc(x_75); lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_69); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; +lean_dec(x_70); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_50, 0); -lean_inc(x_76); -lean_dec(x_50); -x_77 = lean_array_fget(x_11, x_76); +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_51, 0); +lean_inc(x_77); +lean_dec(x_51); +x_78 = lean_array_fget(x_12, x_77); +lean_inc(x_23); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); @@ -14975,30 +14977,30 @@ lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_12); +lean_inc(x_13); +lean_inc(x_9); lean_inc(x_7); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_78 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12, x_34, x_77, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -if (lean_obj_tag(x_78) == 0) +x_79 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13, x_35, x_78, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_78, 1); -lean_inc(x_79); -lean_dec(x_78); -x_80 = l_Array_eraseIdx___rarg(x_11, x_76, lean_box(0)); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_80); -x_27 = x_81; -x_28 = x_79; -goto block_33; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = l_Array_eraseIdx___rarg(x_12, x_77, lean_box(0)); +x_82 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_82, 0, x_81); +x_28 = x_82; +x_29 = x_80; +goto block_34; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_76); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_77); +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -15006,91 +15008,90 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_82 = lean_ctor_get(x_78, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_78, 1); +x_83 = lean_ctor_get(x_79, 0); lean_inc(x_83); -if (lean_is_exclusive(x_78)) { - lean_ctor_release(x_78, 0); - lean_ctor_release(x_78, 1); - x_84 = x_78; +x_84 = lean_ctor_get(x_79, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_85 = x_79; } else { - lean_dec_ref(x_78); - x_84 = lean_box(0); + lean_dec_ref(x_79); + x_85 = lean_box(0); } -if (lean_is_scalar(x_84)) { - x_85 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); } else { - x_85 = x_84; + x_86 = x_85; } -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_83); -return x_85; +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; } } } -block_33: +block_34: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_10, 2); -x_31 = lean_nat_add(x_12, x_30); -lean_dec(x_12); -x_11 = x_29; -x_12 = x_31; -x_13 = lean_box(0); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_ctor_get(x_11, 2); +x_32 = lean_nat_add(x_13, x_31); +lean_dec(x_13); +x_12 = x_30; +x_13 = x_32; x_14 = lean_box(0); -x_23 = x_28; +x_15 = lean_box(0); +x_24 = x_29; goto _start; } -block_47: +block_48: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -lean_dec(x_36); -x_37 = l_Lean_MessageData_ofName(x_35); -x_38 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___closed__2; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -x_40 = l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___closed__4; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__1(x_34, x_41, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -lean_dec(x_22); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +lean_dec(x_37); +x_38 = l_Lean_MessageData_ofName(x_36); +x_39 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4___closed__2; +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___closed__4; +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__1(x_35, x_42, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +lean_dec(x_23); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_34); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_dec(x_35); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) { -return x_42; +return x_43; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_ctor_get(x_42, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_43, 0); +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_42); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_dec(x_43); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } @@ -15407,7 +15408,7 @@ x_2 = lean_box_usize(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; @@ -15463,28 +15464,23 @@ lean_inc(x_9); x_30 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac(x_4, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_29); if (lean_obj_tag(x_30) == 0) { -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = lean_ctor_get(x_30, 0); -x_33 = lean_ctor_get(x_30, 1); -x_34 = lean_unsigned_to_nat(0u); -x_35 = lean_nat_dec_lt(x_34, x_5); -if (x_35 == 0) +if (x_5 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -lean_free_object(x_30); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_dec(x_6); -x_36 = l_Lean_Elab_Tactic_getGoals___rarg(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_33); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = l_List_appendTR___rarg(x_37, x_32); -x_40 = l_Lean_Elab_Tactic_setGoals(x_39, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_38); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Elab_Tactic_getGoals___rarg(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_32); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_List_appendTR___rarg(x_34, x_31); +x_37 = l_Lean_Elab_Tactic_setGoals(x_36, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_35); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -15493,33 +15489,39 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_40, 0); -lean_dec(x_42); -x_43 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -lean_ctor_set(x_40, 0, x_43); -return x_40; +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +x_40 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +lean_ctor_set(x_37, 0, x_40); +return x_37; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_40, 1); -lean_inc(x_44); -lean_dec(x_40); -x_45 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_44); -return x_46; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 1); +lean_inc(x_41); +lean_dec(x_37); +x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +return x_43; } } else { -uint8_t x_47; -x_47 = l_List_isEmpty___rarg(x_32); +uint8_t x_44; +x_44 = !lean_is_exclusive(x_30); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = lean_ctor_get(x_30, 0); +x_46 = lean_ctor_get(x_30, 1); +x_47 = l_List_isEmpty___rarg(x_45); if (x_47 == 0) { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; @@ -15535,11 +15537,11 @@ lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); x_53 = 2; lean_inc(x_15); -x_54 = l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(x_52, x_53, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_33); +x_54 = l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(x_52, x_53, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_46); x_55 = lean_ctor_get(x_54, 1); lean_inc(x_55); lean_dec(x_54); -x_56 = l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(x_32, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_55); +x_56 = l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(x_45, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_55); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -15596,7 +15598,7 @@ return x_66; else { lean_object* x_67; -lean_dec(x_32); +lean_dec(x_45); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -15611,135 +15613,89 @@ lean_ctor_set(x_30, 0, x_67); return x_30; } } -} else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_object* x_68; lean_object* x_69; uint8_t x_70; x_68 = lean_ctor_get(x_30, 0); x_69 = lean_ctor_get(x_30, 1); lean_inc(x_69); lean_inc(x_68); lean_dec(x_30); -x_70 = lean_unsigned_to_nat(0u); -x_71 = lean_nat_dec_lt(x_70, x_5); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -lean_dec(x_6); -x_72 = l_Lean_Elab_Tactic_getGoals___rarg(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_69); -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); -lean_inc(x_74); -lean_dec(x_72); -x_75 = l_List_appendTR___rarg(x_73, x_68); -x_76 = l_Lean_Elab_Tactic_setGoals(x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_74); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_78 = x_76; -} else { - lean_dec_ref(x_76); - x_78 = lean_box(0); -} -x_79 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -if (lean_is_scalar(x_78)) { - x_80 = lean_alloc_ctor(0, 2, 0); -} else { - x_80 = x_78; -} -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_77); -return x_80; -} -else -{ -uint8_t x_81; -x_81 = l_List_isEmpty___rarg(x_68); -if (x_81 == 0) +x_70 = l_List_isEmpty___rarg(x_68); +if (x_70 == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_82 = l_Lean_MessageData_ofName(x_6); -x_83 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6___closed__2; -x_84 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_82); -x_85 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__3; -x_86 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_86, 0, x_84); -lean_ctor_set(x_86, 1, x_85); -x_87 = 2; +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_71 = l_Lean_MessageData_ofName(x_6); +x_72 = l_Lean_Elab_Tactic_ElimApp_evalAlts_applyAltStx___lambda__6___closed__2; +x_73 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__3; +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_76 = 2; lean_inc(x_15); -x_88 = l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(x_86, x_87, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_69); -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -lean_dec(x_88); -x_90 = l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(x_68, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_89); +x_77 = l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(x_75, x_76, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_69); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_79 = l_List_forM___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__7(x_68, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_78); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -if (lean_obj_tag(x_90) == 0) +if (lean_obj_tag(x_79) == 0) { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_92 = x_90; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_81 = x_79; } else { - lean_dec_ref(x_90); - x_92 = lean_box(0); + lean_dec_ref(x_79); + x_81 = lean_box(0); } -x_93 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -if (lean_is_scalar(x_92)) { - x_94 = lean_alloc_ctor(0, 2, 0); +x_82 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +if (lean_is_scalar(x_81)) { + x_83 = lean_alloc_ctor(0, 2, 0); } else { - x_94 = x_92; + x_83 = x_81; } -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_91); -return x_94; +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_80); +return x_83; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_95 = lean_ctor_get(x_90, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_90, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_97 = x_90; +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_ctor_get(x_79, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_79, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_86 = x_79; } else { - lean_dec_ref(x_90); - x_97 = lean_box(0); + lean_dec_ref(x_79); + x_86 = lean_box(0); } -if (lean_is_scalar(x_97)) { - x_98 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(1, 2, 0); } else { - x_98 = x_97; + x_87 = x_86; } -lean_ctor_set(x_98, 0, x_95); -lean_ctor_set(x_98, 1, x_96); -return x_98; +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_85); +return x_87; } } else { -lean_object* x_99; lean_object* x_100; +lean_object* x_88; lean_object* x_89; lean_dec(x_68); lean_dec(x_16); lean_dec(x_15); @@ -15750,18 +15706,18 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_6); -x_99 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_69); -return x_100; +x_88 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_69); +return x_89; } } } } else { -uint8_t x_101; +uint8_t x_90; lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -15771,29 +15727,29 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_6); -x_101 = !lean_is_exclusive(x_30); -if (x_101 == 0) +x_90 = !lean_is_exclusive(x_30); +if (x_90 == 0) { return x_30; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_30, 0); -x_103 = lean_ctor_get(x_30, 1); -lean_inc(x_103); -lean_inc(x_102); +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_30, 0); +x_92 = lean_ctor_get(x_30, 1); +lean_inc(x_92); +lean_inc(x_91); lean_dec(x_30); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -return x_104; +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; } } } else { -uint8_t x_105; +uint8_t x_94; lean_dec(x_22); lean_dec(x_16); lean_dec(x_15); @@ -15804,29 +15760,29 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_6); -x_105 = !lean_is_exclusive(x_28); -if (x_105 == 0) +x_94 = !lean_is_exclusive(x_28); +if (x_94 == 0) { return x_28; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_28, 0); -x_107 = lean_ctor_get(x_28, 1); -lean_inc(x_107); -lean_inc(x_106); +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_28, 0); +x_96 = lean_ctor_get(x_28, 1); +lean_inc(x_96); +lean_inc(x_95); lean_dec(x_28); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; } } } else { -uint8_t x_109; +uint8_t x_98; lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -15838,36 +15794,35 @@ lean_dec(x_9); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_109 = !lean_is_exclusive(x_21); -if (x_109 == 0) +x_98 = !lean_is_exclusive(x_21); +if (x_98 == 0) { return x_21; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_21, 0); -x_111 = lean_ctor_get(x_21, 1); -lean_inc(x_111); -lean_inc(x_110); +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_21, 0); +x_100 = lean_ctor_get(x_21, 1); +lean_inc(x_100); +lean_inc(x_99); lean_dec(x_21); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, size_t x_12, size_t x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, size_t x_11, size_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22) { _start: { -uint8_t x_24; -x_24 = lean_usize_dec_lt(x_13, x_12); -if (x_24 == 0) +uint8_t x_23; +x_23 = lean_usize_dec_lt(x_12, x_11); +if (x_23 == 0) { -lean_object* x_25; -lean_dec(x_22); +lean_object* x_24; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -15875,103 +15830,103 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_23); -return x_25; +lean_dec(x_3); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_13); +lean_ctor_set(x_24, 1, x_22); +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_14); -x_26 = lean_array_uget(x_11, x_13); -x_36 = lean_ctor_get(x_26, 0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_13); +x_25 = lean_array_uget(x_10, x_12); +x_35 = lean_ctor_get(x_25, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_25, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_26, 1); +x_37 = lean_ctor_get(x_25, 2); lean_inc(x_37); -x_38 = lean_ctor_get(x_26, 2); -lean_inc(x_38); -lean_dec(x_26); -lean_inc(x_17); -lean_inc(x_36); -x_39 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields(x_1, x_36, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -if (lean_obj_tag(x_39) == 0) +lean_dec(x_25); +lean_inc(x_16); +lean_inc(x_35); +x_38 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields(x_1, x_35, x_16, x_17, x_18, x_19, x_20, x_21, x_22); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; -x_40 = lean_ctor_get(x_39, 0); +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_box(0); -x_43 = 0; -lean_inc(x_22); +lean_dec(x_38); +x_41 = lean_box(0); +x_42 = 0; lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -x_44 = l_Lean_Meta_introNCore(x_38, x_40, x_42, x_43, x_43, x_19, x_20, x_21, x_22, x_41); -if (lean_obj_tag(x_44) == 0) +lean_inc(x_18); +x_43 = l_Lean_Meta_introNCore(x_37, x_39, x_41, x_42, x_42, x_18, x_19, x_20, x_21, x_40); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_45 = lean_ctor_get(x_44, 0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); lean_inc(x_45); +lean_dec(x_43); x_46 = lean_ctor_get(x_44, 1); lean_inc(x_46); lean_dec(x_44); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); +x_47 = lean_box(0); x_48 = lean_box(0); -x_49 = lean_box(0); -lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -lean_inc(x_47); -lean_inc(x_4); -x_50 = l_Lean_Meta_Cases_unifyEqs_x3f(x_4, x_47, x_48, x_49, x_19, x_20, x_21, x_22, x_46); -if (lean_obj_tag(x_50) == 0) +lean_inc(x_18); +lean_inc(x_46); +lean_inc(x_3); +x_49 = l_Lean_Meta_Cases_unifyEqs_x3f(x_3, x_46, x_47, x_48, x_18, x_19, x_20, x_21, x_45); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_51; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -if (lean_obj_tag(x_51) == 0) +lean_object* x_50; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_52; lean_object* x_53; -lean_dec(x_47); -lean_dec(x_37); +lean_object* x_51; lean_object* x_52; +lean_dec(x_46); lean_dec(x_36); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); -x_53 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; -x_27 = x_53; -x_28 = x_52; -goto block_35; +lean_dec(x_35); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2___closed__1; +x_26 = x_52; +x_27 = x_51; +goto block_34; } else { -lean_object* x_54; uint8_t x_55; -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_ctor_get_uint8(x_37, sizeof(void*)*3); -lean_dec(x_37); -if (x_55 == 0) +lean_object* x_53; uint8_t x_54; +x_53 = lean_ctor_get(x_50, 0); +lean_inc(x_53); +lean_dec(x_50); +x_54 = lean_ctor_get_uint8(x_36, sizeof(void*)*3); +lean_dec(x_36); +if (x_54 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_50, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_49, 1); +lean_inc(x_55); +lean_dec(x_49); +x_56 = lean_ctor_get(x_53, 0); lean_inc(x_56); -lean_dec(x_50); -x_57 = lean_ctor_get(x_54, 0); +x_57 = lean_ctor_get(x_53, 1); lean_inc(x_57); -x_58 = lean_ctor_get(x_54, 1); -lean_inc(x_58); -lean_dec(x_54); -lean_inc(x_22); +lean_dec(x_53); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); @@ -15979,25 +15934,25 @@ lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_7); -x_59 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_6, x_7, x_58, x_2, x_8, x_36, x_47, x_57, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_56); -lean_dec(x_47); -if (lean_obj_tag(x_59) == 0) +lean_inc(x_14); +lean_inc(x_6); +x_58 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_5, x_6, x_57, x_2, x_7, x_35, x_46, x_56, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_55); +lean_dec(x_46); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_60; lean_object* x_61; -x_60 = lean_ctor_get(x_59, 0); +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); +lean_dec(x_58); +x_26 = x_59; x_27 = x_60; -x_28 = x_61; -goto block_35; +goto block_34; } else { -uint8_t x_62; -lean_dec(x_22); +uint8_t x_61; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16005,60 +15960,60 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_62 = !lean_is_exclusive(x_59); -if (x_62 == 0) +lean_dec(x_3); +x_61 = !lean_is_exclusive(x_58); +if (x_61 == 0) { -return x_59; +return x_58; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_59, 0); -x_64 = lean_ctor_get(x_59, 1); -lean_inc(x_64); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_58, 0); +x_63 = lean_ctor_get(x_58, 1); lean_inc(x_63); -lean_dec(x_59); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_inc(x_62); +lean_dec(x_58); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; -lean_dec(x_47); -x_66 = lean_ctor_get(x_50, 1); +lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; +lean_dec(x_46); +x_65 = lean_ctor_get(x_49, 1); +lean_inc(x_65); +lean_dec(x_49); +x_66 = lean_ctor_get(x_53, 0); lean_inc(x_66); -lean_dec(x_50); -x_67 = lean_ctor_get(x_54, 0); +x_67 = lean_ctor_get(x_53, 1); lean_inc(x_67); -x_68 = lean_ctor_get(x_54, 1); -lean_inc(x_68); -lean_dec(x_54); -x_69 = 1; -lean_inc(x_22); +lean_dec(x_53); +x_68 = 1; lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -lean_inc(x_5); -x_70 = l_Lean_Meta_introNCore(x_67, x_5, x_42, x_43, x_69, x_19, x_20, x_21, x_22, x_66); -if (lean_obj_tag(x_70) == 0) +lean_inc(x_18); +lean_inc(x_4); +x_69 = l_Lean_Meta_introNCore(x_66, x_4, x_41, x_42, x_68, x_18, x_19, x_20, x_21, x_65); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_71 = lean_ctor_get(x_70, 0); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); lean_inc(x_71); +lean_dec(x_69); x_72 = lean_ctor_get(x_70, 1); lean_inc(x_72); lean_dec(x_70); -x_73 = lean_ctor_get(x_71, 1); -lean_inc(x_73); -lean_dec(x_71); -lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); @@ -16066,26 +16021,26 @@ lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_73); -lean_inc(x_7); -x_74 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_6, x_7, x_68, x_2, x_8, x_36, x_73, x_73, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_72); -lean_dec(x_73); -if (lean_obj_tag(x_74) == 0) +lean_inc(x_14); +lean_inc(x_72); +lean_inc(x_6); +x_73 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_5, x_6, x_67, x_2, x_7, x_35, x_72, x_72, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_71); +lean_dec(x_72); +if (lean_obj_tag(x_73) == 0) { -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_74, 0); +lean_object* x_74; lean_object* x_75; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); -lean_inc(x_76); -lean_dec(x_74); +lean_dec(x_73); +x_26 = x_74; x_27 = x_75; -x_28 = x_76; -goto block_35; +goto block_34; } else { -uint8_t x_77; -lean_dec(x_22); +uint8_t x_76; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16093,35 +16048,35 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_77 = !lean_is_exclusive(x_74); -if (x_77 == 0) +lean_dec(x_3); +x_76 = !lean_is_exclusive(x_73); +if (x_76 == 0) { -return x_74; +return x_73; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_74, 0); -x_79 = lean_ctor_get(x_74, 1); -lean_inc(x_79); +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_73, 0); +x_78 = lean_ctor_get(x_73, 1); lean_inc(x_78); -lean_dec(x_74); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; +lean_inc(x_77); +lean_dec(x_73); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; } } } else { -uint8_t x_81; -lean_dec(x_68); -lean_dec(x_36); -lean_dec(x_22); +uint8_t x_80; +lean_dec(x_67); +lean_dec(x_35); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16129,26 +16084,27 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_81 = !lean_is_exclusive(x_70); -if (x_81 == 0) +lean_dec(x_3); +x_80 = !lean_is_exclusive(x_69); +if (x_80 == 0) { -return x_70; +return x_69; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_70, 0); -x_83 = lean_ctor_get(x_70, 1); -lean_inc(x_83); +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_69, 0); +x_82 = lean_ctor_get(x_69, 1); lean_inc(x_82); -lean_dec(x_70); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -return x_84; +lean_inc(x_81); +lean_dec(x_69); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; } } } @@ -16156,11 +16112,10 @@ return x_84; } else { -uint8_t x_85; -lean_dec(x_47); -lean_dec(x_37); +uint8_t x_84; +lean_dec(x_46); lean_dec(x_36); -lean_dec(x_22); +lean_dec(x_35); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16168,35 +16123,35 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_85 = !lean_is_exclusive(x_50); -if (x_85 == 0) +lean_dec(x_3); +x_84 = !lean_is_exclusive(x_49); +if (x_84 == 0) { -return x_50; +return x_49; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_50, 0); -x_87 = lean_ctor_get(x_50, 1); -lean_inc(x_87); +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_49, 0); +x_86 = lean_ctor_get(x_49, 1); lean_inc(x_86); -lean_dec(x_50); -x_88 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -return x_88; +lean_inc(x_85); +lean_dec(x_49); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; } } } else { -uint8_t x_89; -lean_dec(x_37); +uint8_t x_88; lean_dec(x_36); -lean_dec(x_22); +lean_dec(x_35); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16204,36 +16159,36 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_89 = !lean_is_exclusive(x_44); -if (x_89 == 0) +lean_dec(x_3); +x_88 = !lean_is_exclusive(x_43); +if (x_88 == 0) { -return x_44; +return x_43; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_44, 0); -x_91 = lean_ctor_get(x_44, 1); -lean_inc(x_91); +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_43, 0); +x_90 = lean_ctor_get(x_43, 1); lean_inc(x_90); -lean_dec(x_44); -x_92 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -return x_92; +lean_inc(x_89); +lean_dec(x_43); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; } } } else { -uint8_t x_93; -lean_dec(x_38); +uint8_t x_92; lean_dec(x_37); lean_dec(x_36); -lean_dec(x_22); +lean_dec(x_35); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16241,34 +16196,34 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_93 = !lean_is_exclusive(x_39); -if (x_93 == 0) +lean_dec(x_3); +x_92 = !lean_is_exclusive(x_38); +if (x_92 == 0) { -return x_39; +return x_38; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_39, 0); -x_95 = lean_ctor_get(x_39, 1); -lean_inc(x_95); +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_38, 0); +x_94 = lean_ctor_get(x_38, 1); lean_inc(x_94); -lean_dec(x_39); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -return x_96; +lean_inc(x_93); +lean_dec(x_38); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } -block_35: +block_34: { -if (lean_obj_tag(x_27) == 0) +if (lean_obj_tag(x_26) == 0) { -lean_object* x_29; lean_object* x_30; -lean_dec(x_22); +lean_object* x_28; lean_object* x_29; lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); @@ -16276,28 +16231,29 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; +lean_dec(x_3); +x_28 = lean_ctor_get(x_26, 0); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } else { -lean_object* x_31; size_t x_32; size_t x_33; -x_31 = lean_ctor_get(x_27, 0); -lean_inc(x_31); -lean_dec(x_27); -x_32 = 1; -x_33 = lean_usize_add(x_13, x_32); -x_13 = x_33; -x_14 = x_31; -x_23 = x_28; +lean_object* x_30; size_t x_31; size_t x_32; +x_30 = lean_ctor_get(x_26, 0); +lean_inc(x_30); +lean_dec(x_26); +x_31 = 1; +x_32 = lean_usize_add(x_12, x_31); +x_12 = x_32; +x_13 = x_30; +x_22 = x_27; goto _start; } } @@ -16307,23 +16263,57 @@ goto _start; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_object* x_19; +uint8_t x_19; +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_55; +x_55 = 0; +x_19 = x_55; +goto block_54; +} +else +{ +uint8_t x_56; +x_56 = 1; +x_19 = x_56; +goto block_54; +} +block_54: +{ +lean_object* x_20; +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_52; +x_52 = l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__1; +x_20 = x_52; +goto block_51; +} +else +{ +lean_object* x_53; +x_53 = lean_ctor_get(x_4, 0); +lean_inc(x_53); +x_20 = x_53; +goto block_51; +} +block_51: +{ +lean_object* x_21; lean_inc(x_16); -x_19 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames(x_2, x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_19) == 0) +x_21 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_checkAltNames(x_2, x_20, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_array_get_size(x_4); -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_unsigned_to_nat(1u); -lean_inc(x_21); -x_24 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_23); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_array_get_size(x_20); +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_23); +lean_ctor_set(x_26, 2, x_25); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -16335,76 +16325,73 @@ lean_inc(x_10); lean_inc(x_8); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_24, x_24, x_2, x_22, lean_box(0), lean_box(0), x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_20); -lean_dec(x_24); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_box(0); -x_29 = lean_array_size(x_26); -x_30 = 0; -x_31 = lean_box(0); -x_32 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_21, x_26, x_28, x_26, x_29, x_30, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_27); +x_27 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20, x_26, x_26, x_2, x_24, lean_box(0), lean_box(0), x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_22); lean_dec(x_26); -lean_dec(x_21); lean_dec(x_4); -if (lean_obj_tag(x_32) == 0) +if (lean_obj_tag(x_27) == 0) { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; lean_object* x_34; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_box(0); +x_31 = lean_array_size(x_28); +x_32 = 0; +x_33 = lean_box(0); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(x_1, x_3, x_5, x_6, x_7, x_8, x_19, x_28, x_30, x_28, x_31, x_32, x_33, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_29); +lean_dec(x_28); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_34; -x_34 = lean_ctor_get(x_32, 0); -lean_dec(x_34); -lean_ctor_set(x_32, 0, x_31); -return x_32; +uint8_t x_35; +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_34, 0); +lean_dec(x_36); +lean_ctor_set(x_34, 0, x_33); +return x_34; } else { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_31); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } else { -uint8_t x_37; -x_37 = !lean_is_exclusive(x_32); -if (x_37 == 0) +uint8_t x_39; +x_39 = !lean_is_exclusive(x_34); +if (x_39 == 0) { -return x_32; +return x_34; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_32, 0); -x_39 = lean_ctor_get(x_32, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_32); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_34, 0); +x_41 = lean_ctor_get(x_34, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_34); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } else { -uint8_t x_41; -lean_dec(x_21); +uint8_t x_43; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16416,30 +16403,30 @@ lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_41 = !lean_is_exclusive(x_25); -if (x_41 == 0) +x_43 = !lean_is_exclusive(x_27); +if (x_43 == 0) { -return x_25; +return x_27; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_25, 0); -x_43 = lean_ctor_get(x_25, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_25); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_27, 0); +x_45 = lean_ctor_get(x_27, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_27); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } else { -uint8_t x_45; +uint8_t x_47; +lean_dec(x_20); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16453,23 +16440,25 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_45 = !lean_is_exclusive(x_19); -if (x_45 == 0) +x_47 = !lean_is_exclusive(x_21); +if (x_47 == 0) { -return x_19; +return x_21; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_19, 0); -x_47 = lean_ctor_get(x_19, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_19); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_21, 0); +x_49 = lean_ctor_get(x_21, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_21); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} } } } @@ -16543,8 +16532,8 @@ x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWit lean_dec(x_13); lean_dec(x_12); lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_28; @@ -16585,17 +16574,19 @@ lean_object* x_20 = _args[19]; lean_object* x_21 = _args[20]; lean_object* x_22 = _args[21]; lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; _start: { -lean_object* x_24; -x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +lean_object* x_25; +x_25 = l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_24; +return x_25; } } LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { @@ -16684,13 +16675,14 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; _start: { -lean_object* x_18; -x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_7); +uint8_t x_18; lean_object* x_19; +x_18 = lean_unbox(x_5); lean_dec(x_5); +x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___lambda__2(x_1, x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_7); lean_dec(x_4); lean_dec(x_1); -return x_18; +return x_19; } } LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8___boxed(lean_object** _args) { @@ -16716,21 +16708,20 @@ lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; lean_object* x_21 = _args[20]; lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; _start: { -size_t x_24; size_t x_25; lean_object* x_26; -x_24 = lean_unbox_usize(x_12); -lean_dec(x_12); -x_25 = lean_unbox_usize(x_13); -lean_dec(x_13); -x_26 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24, x_25, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +uint8_t x_23; size_t x_24; size_t x_25; lean_object* x_26; +x_23 = lean_unbox(x_7); +lean_dec(x_7); +x_24 = lean_unbox_usize(x_11); lean_dec(x_11); +x_25 = lean_unbox_usize(x_12); +lean_dec(x_12); +x_26 = l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_23, x_8, x_9, x_10, x_24, x_25, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_3); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_26; @@ -17393,162 +17384,163 @@ lean_ctor_set_uint8(x_6, sizeof(void*)*4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { _start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_21 = lean_box(0); -x_22 = lean_box(2); -x_23 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_22 = lean_box(0); +x_23 = lean_box(2); +x_24 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; lean_inc(x_1); -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_1); +x_25 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_1); lean_inc(x_2); -x_25 = lean_io_promise_result(x_2); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_21); -lean_ctor_set(x_26, 1, x_25); -x_27 = lean_unsigned_to_nat(0u); -x_28 = l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__1; -x_29 = l_Array_zipWithAux___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__1(x_1, x_11, x_27, x_28); -x_30 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__11; -x_31 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_24); -lean_ctor_set(x_31, 2, x_21); -lean_ctor_set(x_31, 3, x_26); -lean_ctor_set(x_31, 4, x_29); -x_32 = lean_ctor_get(x_3, 1); -lean_inc(x_32); -x_33 = lean_io_promise_resolve(x_31, x_32, x_20); -lean_dec(x_32); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); +x_26 = lean_io_promise_result(x_2); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_22); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__1; +x_30 = l_Array_zipWithAux___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__1(x_1, x_12, x_28, x_29); +lean_dec(x_1); +x_31 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__11; +x_32 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_25); +lean_ctor_set(x_32, 2, x_22); +lean_ctor_set(x_32, 3, x_27); +lean_ctor_set(x_32, 4, x_30); +x_33 = lean_ctor_get(x_3, 1); +lean_inc(x_33); +x_34 = lean_io_promise_resolve(x_32, x_33, x_21); lean_dec(x_33); -x_35 = lean_array_get_size(x_11); -x_36 = lean_mk_empty_array_with_capacity(x_35); -x_37 = l_Array_mapFinIdxM_map___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__2(x_3, x_11, x_11, x_35, x_27, lean_box(0), x_36); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_array_get_size(x_12); +x_37 = lean_mk_empty_array_with_capacity(x_36); +x_38 = l_Array_mapFinIdxM_map___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__2(x_3, x_12, x_12, x_36, x_28, lean_box(0), x_37); +lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -x_38 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental(x_4, x_5, x_6, x_1, x_7, x_8, x_9, x_10, x_37, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_34); -lean_dec(x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); +x_39 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithIncremental(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_35); lean_dec(x_38); -x_40 = l_Lean_Elab_Tactic_saveState___rarg(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = l_Lean_Elab_Tactic_saveState___rarg(x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_40); +lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_42 = lean_ctor_get(x_40, 0); -x_43 = lean_ctor_get(x_40, 1); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_40, 1, x_44); -lean_ctor_set(x_40, 0, x_30); -x_45 = lean_io_promise_resolve(x_40, x_2, x_43); +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_43 = lean_ctor_get(x_41, 0); +x_44 = lean_ctor_get(x_41, 1); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_41, 1, x_45); +lean_ctor_set(x_41, 0, x_31); +x_46 = lean_io_promise_resolve(x_41, x_2, x_44); lean_dec(x_2); -x_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) +x_47 = !lean_is_exclusive(x_46); +if (x_47 == 0) { -return x_45; +return x_46; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_45, 0); -x_48 = lean_ctor_get(x_45, 1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_46, 0); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_45); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_dec(x_46); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; } } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_50 = lean_ctor_get(x_40, 0); -x_51 = lean_ctor_get(x_40, 1); +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_51 = lean_ctor_get(x_41, 0); +x_52 = lean_ctor_get(x_41, 1); +lean_inc(x_52); lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_40); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_50); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_30); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_io_promise_resolve(x_53, x_2, x_51); +lean_dec(x_41); +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_51); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_31); +lean_ctor_set(x_54, 1, x_53); +x_55 = lean_io_promise_resolve(x_54, x_2, x_52); lean_dec(x_2); -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); +x_56 = lean_ctor_get(x_55, 0); lean_inc(x_56); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - x_57 = x_54; +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_58 = x_55; } else { - lean_dec_ref(x_54); - x_57 = lean_box(0); + lean_dec_ref(x_55); + x_58 = lean_box(0); } -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(0, 2, 0); } else { - x_58 = x_57; + x_59 = x_58; } -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; } } else { -uint8_t x_59; +uint8_t x_60; +lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); lean_dec(x_2); -x_59 = !lean_is_exclusive(x_38); -if (x_59 == 0) +x_60 = !lean_is_exclusive(x_39); +if (x_60 == 0) { -return x_38; +return x_39; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_38, 0); -x_61 = lean_ctor_get(x_38, 1); +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_39, 0); +x_62 = lean_ctor_get(x_39, 1); +lean_inc(x_62); lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_38); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_dec(x_39); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; } } } @@ -17556,122 +17548,120 @@ return x_62; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___boxed), 20, 10); -lean_closure_set(x_21, 0, x_1); -lean_closure_set(x_21, 1, x_11); -lean_closure_set(x_21, 2, x_2); -lean_closure_set(x_21, 3, x_3); -lean_closure_set(x_21, 4, x_4); -lean_closure_set(x_21, 5, x_5); -lean_closure_set(x_21, 6, x_6); -lean_closure_set(x_21, 7, x_7); -lean_closure_set(x_21, 8, x_8); -lean_closure_set(x_21, 9, x_9); -x_22 = l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; -x_23 = l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__3(x_22, x_10, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -return x_23; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_array_get_size(x_1); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___boxed), 21, 11); +lean_closure_set(x_22, 0, x_1); +lean_closure_set(x_22, 1, x_11); +lean_closure_set(x_22, 2, x_2); +lean_closure_set(x_22, 3, x_3); +lean_closure_set(x_22, 4, x_4); +lean_closure_set(x_22, 5, x_5); +lean_closure_set(x_22, 6, x_6); +lean_closure_set(x_22, 7, x_7); +lean_closure_set(x_22, 8, x_8); +lean_closure_set(x_22, 9, x_9); +lean_closure_set(x_22, 10, x_10); +x_23 = l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; +x_24 = l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__3(x_23, x_21, x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_24; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = lean_array_get_size(x_4); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) +if (lean_obj_tag(x_4) == 0) { -lean_object* x_21; lean_object* x_22; -lean_dec(x_18); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_21, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_7); lean_dec(x_3); lean_dec(x_1); -return x_22; +return x_19; } else { -lean_object* x_23; -x_23 = lean_ctor_get(x_11, 8); -lean_inc(x_23); -if (lean_obj_tag(x_23) == 0) +lean_object* x_20; +x_20 = lean_ctor_get(x_11, 8); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_18); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_24, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_21, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_7); lean_dec(x_3); lean_dec(x_1); -return x_25; +return x_22; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_dec(x_23); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3___boxed), 20, 10); -lean_closure_set(x_27, 0, x_4); -lean_closure_set(x_27, 1, x_26); -lean_closure_set(x_27, 2, x_1); -lean_closure_set(x_27, 3, x_2); -lean_closure_set(x_27, 4, x_3); -lean_closure_set(x_27, 5, x_5); -lean_closure_set(x_27, 6, x_6); -lean_closure_set(x_27, 7, x_7); -lean_closure_set(x_27, 8, x_8); -lean_closure_set(x_27, 9, x_18); -x_28 = l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot; -x_29 = l_Lean_Language_withAlwaysResolvedPromise___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__7(x_28, x_27, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_29) == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_4, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_20, 0); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3___boxed), 20, 10); +lean_closure_set(x_25, 0, x_23); +lean_closure_set(x_25, 1, x_24); +lean_closure_set(x_25, 2, x_1); +lean_closure_set(x_25, 3, x_2); +lean_closure_set(x_25, 4, x_3); +lean_closure_set(x_25, 5, x_4); +lean_closure_set(x_25, 6, x_5); +lean_closure_set(x_25, 7, x_6); +lean_closure_set(x_25, 8, x_7); +lean_closure_set(x_25, 9, x_8); +x_26 = l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot; +x_27 = l_Lean_Language_withAlwaysResolvedPromise___at_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___spec__7(x_26, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_27) == 0) { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_29, 0); -lean_dec(x_31); -x_32 = lean_box(0); -lean_ctor_set(x_29, 0, x_32); -return x_29; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); +return x_27; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_29, 1); -lean_inc(x_33); -lean_dec(x_29); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -return x_35; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } else { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_29); -if (x_36 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_27); +if (x_34 == 0) { -return x_29; +return x_27; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_29, 0); -x_38 = lean_ctor_get(x_29, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_29); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_27, 0); +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_27); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } @@ -17815,15 +17805,16 @@ lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; _start: { -lean_object* x_21; -x_21 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_11); -lean_dec(x_9); +lean_object* x_22; +x_22 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_6); lean_dec(x_4); -return x_21; +return x_22; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__3___boxed(lean_object** _args) { @@ -18002,41 +17993,36 @@ return x_13; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_array_get_size(x_4); -x_20 = lean_unsigned_to_nat(0u); -x_21 = lean_nat_dec_lt(x_20, x_19); -lean_dec(x_19); -if (x_21 == 0) +if (lean_obj_tag(x_4) == 0) { -lean_object* x_22; +lean_object* x_19; lean_dec(x_5); -x_22 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo(x_1, x_2, x_3, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -return x_22; +x_19 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo(x_1, x_2, x_3, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +return x_19; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_23 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg(x_17, x_18); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1___boxed), 18, 9); -lean_closure_set(x_26, 0, x_24); -lean_closure_set(x_26, 1, x_1); -lean_closure_set(x_26, 2, x_2); -lean_closure_set(x_26, 3, x_3); -lean_closure_set(x_26, 4, x_4); -lean_closure_set(x_26, 5, x_6); -lean_closure_set(x_26, 6, x_7); -lean_closure_set(x_26, 7, x_8); -lean_closure_set(x_26, 8, x_9); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__2___boxed), 11, 1); -lean_closure_set(x_27, 0, x_5); -x_28 = l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_evalTactic_eval___spec__3(x_26, x_27, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_25); -return x_28; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_20 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg(x_17, x_18); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1___boxed), 18, 9); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_1); +lean_closure_set(x_23, 2, x_2); +lean_closure_set(x_23, 3, x_3); +lean_closure_set(x_23, 4, x_4); +lean_closure_set(x_23, 5, x_6); +lean_closure_set(x_23, 6, x_7); +lean_closure_set(x_23, 7, x_8); +lean_closure_set(x_23, 8, x_9); +x_24 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__2___boxed), 11, 1); +lean_closure_set(x_24, 0, x_5); +x_25 = l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_evalTactic_eval___spec__3(x_23, x_24, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_22); +return x_25; } } } @@ -20687,42 +20673,101 @@ uint8_t x_2; x_2 = l_Lean_Syntax_isNone(x_1); if (x_2 == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_3 = lean_unsigned_to_nat(0u); x_4 = l_Lean_Syntax_getArg(x_1, x_3); -x_5 = l_Lean_Syntax_getArgs(x_4); -x_6 = lean_unsigned_to_nat(2u); -x_7 = l_Array_toSubarray___rarg(x_5, x_3, x_6); -x_8 = l_Array_ofSubarray___rarg(x_7); +x_5 = lean_unsigned_to_nat(2u); +x_6 = l_Lean_Syntax_getArg(x_4, x_5); +x_7 = l_Lean_Syntax_getNumArgs(x_6); +x_8 = lean_nat_dec_lt(x_3, x_7); lean_dec(x_7); -x_9 = lean_box(2); -x_10 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; -x_11 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_10); -lean_ctor_set(x_11, 2, x_8); -x_12 = l_Lean_Syntax_getArg(x_4, x_6); +x_9 = l_Lean_Syntax_getArgs(x_4); +x_10 = l_Array_toSubarray___rarg(x_9, x_3, x_5); +x_11 = l_Array_ofSubarray___rarg(x_10); +lean_dec(x_10); +x_12 = lean_box(2); +x_13 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; +x_14 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +lean_ctor_set(x_14, 2, x_11); +if (x_8 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_6); +x_15 = l_Lean_Syntax_getArg(x_4, x_3); lean_dec(x_4); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -return x_13; +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } else { -lean_object* x_14; -x_14 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2; -return x_14; +lean_object* x_17; +lean_dec(x_4); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_14); +lean_ctor_set(x_17, 1, x_6); +return x_17; +} +} +else +{ +lean_object* x_18; +x_18 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2; +return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1() { _start: { -lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_Syntax_getArgs(x_2); -x_13 = lean_apply_10(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_13; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_ElimApp_mkElimApp___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = l_Lean_Syntax_isNone(x_1); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; +lean_inc(x_3); +x_15 = l_Lean_Syntax_isOfKind(x_3, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_3); +x_16 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1; +x_17 = lean_apply_10(x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = l_Lean_Syntax_getArgs(x_3); +lean_dec(x_3); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_apply_10(x_2, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_20; +} +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_3); +x_21 = lean_box(0); +x_22 = lean_apply_10(x_2, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_22; +} } } static lean_object* _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1() { @@ -20737,8 +20782,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tac _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed), 11, 1); -lean_closure_set(x_12, 0, x_2); +lean_inc(x_1); +x_12 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed), 12, 2); +lean_closure_set(x_12, 0, x_1); +lean_closure_set(x_12, 1, x_2); x_13 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1; x_14 = l_Lean_Elab_Term_withNarrowedTacticReuse___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___spec__1___rarg(x_13, x_12, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); return x_14; @@ -20761,13 +20808,13 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; -x_12 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_2); -return x_12; +lean_object* x_13; +x_13 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_1); +return x_13; } } static lean_object* _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getOptPreTacOfOptInductionAlts___closed__1() { @@ -22320,7 +22367,8 @@ x_150 = lean_ctor_get(x_147, 1); x_151 = lean_ctor_get(x_149, 0); lean_inc(x_151); lean_dec(x_149); -x_152 = lean_environment_find(x_151, x_146); +x_152 = l_Lean_Environment_find_x3f(x_151, x_146); +lean_dec(x_146); if (lean_obj_tag(x_152) == 0) { uint8_t x_153; lean_object* x_154; @@ -23196,7 +23244,8 @@ lean_dec(x_147); x_405 = lean_ctor_get(x_403, 0); lean_inc(x_405); lean_dec(x_403); -x_406 = lean_environment_find(x_405, x_146); +x_406 = l_Lean_Environment_find_x3f(x_405, x_146); +lean_dec(x_146); if (lean_obj_tag(x_406) == 0) { uint8_t x_407; lean_object* x_408; @@ -25286,7 +25335,7 @@ lean_dec(x_2); return x_10; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1() { _start: { lean_object* x_1; @@ -25294,7 +25343,7 @@ x_1 = lean_mk_string_unchecked("tactic", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2() { _start: { lean_object* x_1; @@ -25302,17 +25351,17 @@ x_1 = lean_mk_string_unchecked("customEliminators", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4() { _start: { lean_object* x_1; @@ -25320,13 +25369,13 @@ x_1 = lean_mk_string_unchecked("enable using custom eliminators in the 'inductio return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 1; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1; -x_3 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1; +x_3 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -25335,26 +25384,26 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___closed__1; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__2; x_3 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__1; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1; -x_5 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1; +x_5 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3; -x_3 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3; +x_3 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -25372,7 +25421,7 @@ x_13 = lean_ctor_get(x_11, 0); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); lean_dec(x_13); -x_15 = lean_environment_find(x_14, x_1); +x_15 = l_Lean_Environment_find_x3f(x_14, x_1); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; @@ -25418,7 +25467,7 @@ lean_dec(x_11); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_environment_find(x_25, x_1); +x_26 = l_Lean_Environment_find_x3f(x_25, x_1); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; lean_object* x_28; lean_object* x_29; @@ -26016,7 +26065,6 @@ lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint x_40 = lean_ctor_get(x_36, 0); x_41 = l_Lean_Name_getPrefix(x_40); lean_dec(x_40); -lean_inc(x_41); x_42 = l_Lean_isInductive___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___spec__1(x_41, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -26063,7 +26111,6 @@ lean_inc(x_50); lean_dec(x_36); x_51 = l_Lean_Name_getPrefix(x_50); lean_dec(x_50); -lean_inc(x_51); x_52 = l_Lean_isInductive___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___spec__1(x_51, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); x_53 = lean_ctor_get(x_52, 0); lean_inc(x_53); @@ -26282,6 +26329,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_11; } } @@ -30801,7 +30849,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___clos lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__4; x_2 = l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__5; -x_3 = lean_unsigned_to_nat(744u); +x_3 = lean_unsigned_to_nat(752u); x_4 = lean_unsigned_to_nat(8u); x_5 = l_Lean_Elab_Tactic_elabCasesTargets___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -33151,7 +33199,7 @@ x_3 = l_Lean_Elab_addBuiltinIncrementalElab(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -33161,7 +33209,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -33171,27 +33219,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3; x_2 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5() { _start: { lean_object* x_1; @@ -33199,17 +33247,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7() { _start: { lean_object* x_1; @@ -33217,47 +33265,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8; x_2 = l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getAltName___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_ElimApp_evalAlts_applyPreTac___spec__1___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10; x_2 = l_Lean_Elab_Tactic_ElimApp_evalAlts_goWithInfo___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12() { _start: { lean_object* x_1; @@ -33265,17 +33313,17 @@ x_1 = lean_mk_string_unchecked("Induction", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14() { _start: { lean_object* x_1; @@ -33283,33 +33331,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15; -x_2 = lean_unsigned_to_nat(11631u); +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15; +x_2 = lean_unsigned_to_nat(11763u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -33578,6 +33626,8 @@ l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductio lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__1); l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__1___closed__2); +l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___lambda__2___closed__1); l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_withAltsOfOptInductionAlts___rarg___closed__1); l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getOptPreTacOfOptInductionAlts___closed__1 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getOptPreTacOfOptInductionAlts___closed__1(); @@ -33630,19 +33680,19 @@ l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__14 = _init_l_L lean_mark_persistent(l_Lean_Elab_Tactic_getInductiveValFromMajor___lambda__2___closed__14); l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___closed__1 = _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_elabTermForElim___closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__2); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__3); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__4); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__5); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572____closed__6); -if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8572_(lean_io_mk_world()); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__1); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__2); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__3); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__4); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__5); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704____closed__6); +if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_8704_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Tactic_tactic_customEliminators = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Tactic_tactic_customEliminators); @@ -33769,39 +33819,39 @@ lean_dec_ref(res); }if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_evalCases__2(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__2); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__3); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__4); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__5); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__6); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__7); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__8); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__9); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__10); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__11); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__12); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__13); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__14); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__15); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631____closed__16); -if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11631_(lean_io_mk_world()); +}l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__1); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__2); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__3); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__4); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__5); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__6); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__7); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__8); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__9); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__10); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__11); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__12); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__13); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__14); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__15); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763____closed__16); +if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Induction___hyg_11763_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Match.c b/stage0/stdlib/Lean/Elab/Tactic/Match.c index f032f7e1bed7..d0d13fef7e13 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Match.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Match.c @@ -116,7 +116,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalMatch__1___closed__6; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_List_forM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalMatch___closed__4; @@ -2494,7 +2494,8 @@ x_30 = lean_ctor_get(x_27, 1); x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); lean_dec(x_29); -x_32 = lean_environment_main_module(x_14); +x_32 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_33 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_33, 0, x_26); lean_ctor_set(x_33, 1, x_32); @@ -2672,7 +2673,8 @@ lean_dec(x_27); x_81 = lean_ctor_get(x_79, 1); lean_inc(x_81); lean_dec(x_79); -x_82 = lean_environment_main_module(x_14); +x_82 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_83 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_83, 0, x_26); lean_ctor_set(x_83, 1, x_82); diff --git a/stage0/stdlib/Lean/Elab/Tactic/RCases.c b/stage0/stdlib/Lean/Elab/Tactic/RCases.c index 971d3a6bf7fb..c6ed8e3e025f 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/RCases.c +++ b/stage0/stdlib/Lean/Elab/Tactic/RCases.c @@ -157,7 +157,7 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; lean_object* l_ReaderT_pure___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_RCases_RCasesPatt_typed_x3f___boxed(lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Elab_Tactic_RCases_tryClearMany_x27___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_RCases_evalObtain___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_RCases_evalRCases___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*); @@ -11568,7 +11568,8 @@ if (lean_is_exclusive(x_112)) { x_116 = lean_ctor_get(x_113, 0); lean_inc(x_116); lean_dec(x_113); -x_117 = lean_environment_find(x_116, x_111); +x_117 = l_Lean_Environment_find_x3f(x_116, x_111); +lean_dec(x_111); if (lean_obj_tag(x_117) == 0) { lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; @@ -12653,7 +12654,8 @@ if (lean_is_exclusive(x_282)) { x_286 = lean_ctor_get(x_283, 0); lean_inc(x_286); lean_dec(x_283); -x_287 = lean_environment_find(x_286, x_281); +x_287 = l_Lean_Environment_find_x3f(x_286, x_281); +lean_dec(x_281); if (lean_obj_tag(x_287) == 0) { lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; @@ -13896,7 +13898,8 @@ if (lean_is_exclusive(x_112)) { x_116 = lean_ctor_get(x_113, 0); lean_inc(x_116); lean_dec(x_113); -x_117 = lean_environment_find(x_116, x_111); +x_117 = l_Lean_Environment_find_x3f(x_116, x_111); +lean_dec(x_111); if (lean_obj_tag(x_117) == 0) { lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; @@ -14981,7 +14984,8 @@ if (lean_is_exclusive(x_282)) { x_286 = lean_ctor_get(x_283, 0); lean_inc(x_286); lean_dec(x_283); -x_287 = lean_environment_find(x_286, x_281); +x_287 = l_Lean_Environment_find_x3f(x_286, x_281); +lean_dec(x_281); if (lean_obj_tag(x_287) == 0) { lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; @@ -15845,7 +15849,8 @@ if (lean_is_exclusive(x_112)) { x_116 = lean_ctor_get(x_113, 0); lean_inc(x_116); lean_dec(x_113); -x_117 = lean_environment_find(x_116, x_111); +x_117 = l_Lean_Environment_find_x3f(x_116, x_111); +lean_dec(x_111); if (lean_obj_tag(x_117) == 0) { lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; @@ -16930,7 +16935,8 @@ if (lean_is_exclusive(x_282)) { x_286 = lean_ctor_get(x_283, 0); lean_inc(x_286); lean_dec(x_283); -x_287 = lean_environment_find(x_286, x_281); +x_287 = l_Lean_Environment_find_x3f(x_286, x_281); +lean_dec(x_281); if (lean_obj_tag(x_287) == 0) { lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; @@ -17794,7 +17800,8 @@ if (lean_is_exclusive(x_112)) { x_116 = lean_ctor_get(x_113, 0); lean_inc(x_116); lean_dec(x_113); -x_117 = lean_environment_find(x_116, x_111); +x_117 = l_Lean_Environment_find_x3f(x_116, x_111); +lean_dec(x_111); if (lean_obj_tag(x_117) == 0) { lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; @@ -18879,7 +18886,8 @@ if (lean_is_exclusive(x_282)) { x_286 = lean_ctor_get(x_283, 0); lean_inc(x_286); lean_dec(x_283); -x_287 = lean_environment_find(x_286, x_281); +x_287 = l_Lean_Environment_find_x3f(x_286, x_281); +lean_dec(x_281); if (lean_obj_tag(x_287) == 0) { lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; @@ -32374,7 +32382,7 @@ lean_dec(x_6); x_37 = l_Lean_Elab_Tactic_RCases_RCasesPatt_parse(x_36, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_37) == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); x_39 = lean_ctor_get(x_37, 1); @@ -32382,36 +32390,80 @@ lean_inc(x_39); lean_dec(x_37); x_40 = l_Lean_Elab_Tactic_RCases_RCasesPatt_typed_x3f(x_5, x_38, x_7); lean_dec(x_7); -x_41 = l_Lean_Elab_Tactic_RCases_RCasesPatt_name_x3f(x_40); -if (lean_obj_tag(x_41) == 0) +x_41 = l_Lean_Elab_Tactic_RCases_RCasesPatt_ref(x_40); +x_42 = l_Lean_Elab_Tactic_RCases_RCasesPatt_name_x3f(x_40); +x_43 = lean_ctor_get(x_13, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_13, 1); +lean_inc(x_44); +x_45 = lean_ctor_get(x_13, 2); +lean_inc(x_45); +x_46 = lean_ctor_get(x_13, 3); +lean_inc(x_46); +x_47 = lean_ctor_get(x_13, 4); +lean_inc(x_47); +x_48 = lean_ctor_get(x_13, 5); +lean_inc(x_48); +x_49 = lean_ctor_get(x_13, 6); +lean_inc(x_49); +x_50 = lean_ctor_get(x_13, 7); +lean_inc(x_50); +x_51 = lean_ctor_get(x_13, 8); +lean_inc(x_51); +x_52 = lean_ctor_get(x_13, 9); +lean_inc(x_52); +x_53 = lean_ctor_get(x_13, 10); +lean_inc(x_53); +x_54 = lean_ctor_get_uint8(x_13, sizeof(void*)*12); +x_55 = lean_ctor_get(x_13, 11); +lean_inc(x_55); +x_56 = lean_ctor_get_uint8(x_13, sizeof(void*)*12 + 1); +x_57 = l_Lean_replaceRef(x_41, x_48); +lean_dec(x_48); +lean_dec(x_41); +x_58 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_58, 0, x_43); +lean_ctor_set(x_58, 1, x_44); +lean_ctor_set(x_58, 2, x_45); +lean_ctor_set(x_58, 3, x_46); +lean_ctor_set(x_58, 4, x_47); +lean_ctor_set(x_58, 5, x_57); +lean_ctor_set(x_58, 6, x_49); +lean_ctor_set(x_58, 7, x_50); +lean_ctor_set(x_58, 8, x_51); +lean_ctor_set(x_58, 9, x_52); +lean_ctor_set(x_58, 10, x_53); +lean_ctor_set(x_58, 11, x_55); +lean_ctor_set_uint8(x_58, sizeof(void*)*12, x_54); +lean_ctor_set_uint8(x_58, sizeof(void*)*12 + 1, x_56); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_42; lean_object* x_43; -x_42 = l_Lean_Elab_Tactic_RCases_RCasesPatt_instInhabited___closed__2; +lean_object* x_59; lean_object* x_60; +x_59 = l_Lean_Elab_Tactic_RCases_RCasesPatt_instInhabited___closed__2; lean_inc(x_14); -lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_43 = l_Lean_MVarId_intro(x_1, x_42, x_11, x_12, x_13, x_14, x_39); -if (lean_obj_tag(x_43) == 0) +x_60 = l_Lean_MVarId_intro(x_1, x_59, x_11, x_12, x_58, x_14, x_39); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_ctor_get(x_44, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = l_Lean_Expr_fvar___override(x_46); -x_49 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg(x_47, x_2, x_3, x_48, x_4, x_40, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_45); -return x_49; +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = lean_ctor_get(x_61, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_61, 1); +lean_inc(x_64); +lean_dec(x_61); +x_65 = l_Lean_Expr_fvar___override(x_63); +x_66 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg(x_64, x_2, x_3, x_65, x_4, x_40, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_62); +return x_66; } else { -uint8_t x_50; +uint8_t x_67; lean_dec(x_40); lean_dec(x_14); lean_dec(x_13); @@ -32423,57 +32475,56 @@ lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_50 = !lean_is_exclusive(x_43); -if (x_50 == 0) +x_67 = !lean_is_exclusive(x_60); +if (x_67 == 0) { -return x_43; +return x_60; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_43, 0); -x_52 = lean_ctor_get(x_43, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_43); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -return x_53; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_60, 0); +x_69 = lean_ctor_get(x_60, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_60); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; } } } else { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_41, 0); -lean_inc(x_54); -lean_dec(x_41); +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_42, 0); +lean_inc(x_71); +lean_dec(x_42); lean_inc(x_14); -lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_55 = l_Lean_MVarId_intro(x_1, x_54, x_11, x_12, x_13, x_14, x_39); -if (lean_obj_tag(x_55) == 0) +x_72 = l_Lean_MVarId_intro(x_1, x_71, x_11, x_12, x_58, x_14, x_39); +if (lean_obj_tag(x_72) == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_ctor_get(x_56, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -x_60 = l_Lean_Expr_fvar___override(x_58); -x_61 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg(x_59, x_2, x_3, x_60, x_4, x_40, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_57); -return x_61; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = lean_ctor_get(x_73, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +x_77 = l_Lean_Expr_fvar___override(x_75); +x_78 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg(x_76, x_2, x_3, x_77, x_4, x_40, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_74); +return x_78; } else { -uint8_t x_62; +uint8_t x_79; lean_dec(x_40); lean_dec(x_14); lean_dec(x_13); @@ -32485,30 +32536,30 @@ lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_62 = !lean_is_exclusive(x_55); -if (x_62 == 0) +x_79 = !lean_is_exclusive(x_72); +if (x_79 == 0) { -return x_55; +return x_72; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_55, 0); -x_64 = lean_ctor_get(x_55, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_55); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_72, 0); +x_81 = lean_ctor_get(x_72, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_72); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } } } else { -uint8_t x_66; +uint8_t x_83; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -32522,23 +32573,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_66 = !lean_is_exclusive(x_37); -if (x_66 == 0) +x_83 = !lean_is_exclusive(x_37); +if (x_83 == 0) { return x_37; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_37, 0); -x_68 = lean_ctor_get(x_37, 1); -lean_inc(x_68); -lean_inc(x_67); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_37, 0); +x_85 = lean_ctor_get(x_37, 1); +lean_inc(x_85); +lean_inc(x_84); lean_dec(x_37); -x_69 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -return x_69; +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } } } diff --git a/stage0/stdlib/Lean/Elab/Term.c b/stage0/stdlib/Lean/Elab/Term.c index 88384ad81972..c25cdb815768 100644 --- a/stage0/stdlib/Lean/Elab/Term.c +++ b/stage0/stdlib/Lean/Elab/Term.c @@ -379,7 +379,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionT LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Tactic_hashCacheKey____x40_Lean_Elab_Term___hyg_948____boxed(lean_object*); static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_applyResult(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__22___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkConst___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -531,7 +531,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_expand LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostpone___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevM_x3f___at_Lean_Elab_Term_resolveLocalName___spec__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Elab_Term_exposeLevelMVars___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Term_withMacroExpansion___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); @@ -546,7 +545,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerMVarErrorCustomInfo___boxed(le static lean_object* l_Lean_Elab_Term_resolveName___closed__3; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withPushMacroExpansionStack___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Elab_Term_mkTypeMismatchError___closed__1; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_resolveName_x27___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Term_expandDeclId___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -792,6 +790,7 @@ lean_object* l_ReaderT_instMonadLift(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instToSnapshotTreeTacticParsedSnapshot; LEAN_EXPORT lean_object* l_Lean_Elab_Term_LVal_getRef(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_withAutoBoundImplicit_loop___rarg___closed__1; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); @@ -1326,7 +1325,7 @@ static lean_object* l___auto____x40_Lean_Elab_Term___hyg_3095____closed__26; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicit(lean_object*); static lean_object* l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getMVarDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicit_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_transform___at_Lean_Elab_Term_exposeLevelMVars___spec__1___closed__1; @@ -1735,6 +1734,7 @@ static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20935____clos static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___closed__1; static lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringLVal(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_3930_(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -56613,7 +56613,8 @@ lean_dec(x_15); x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); -x_19 = lean_environment_main_module(x_18); +x_19 = l_Lean_Environment_mainModule(x_18); +lean_dec(x_18); x_20 = lean_ctor_get(x_12, 10); lean_inc(x_20); x_21 = l_Lean_addMacroScope(x_19, x_1, x_20); @@ -57311,7 +57312,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -57493,7 +57495,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); @@ -57931,7 +57934,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -58113,7 +58117,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); @@ -71256,8 +71261,7 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -lean_inc(x_1); -x_14 = lean_environment_find(x_13, x_1); +x_14 = l_Lean_Environment_find_x3f(x_13, x_1); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -71298,8 +71302,7 @@ lean_dec(x_9); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -lean_inc(x_1); -x_26 = lean_environment_find(x_25, x_1); +x_26 = l_Lean_Environment_find_x3f(x_25, x_1); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; @@ -74987,7 +74990,7 @@ lean_dec(x_134); x_137 = lean_ctor_get(x_135, 0); lean_inc(x_137); lean_dec(x_135); -x_138 = l_Lean_Kernel_isDiagnosticsEnabled(x_137); +x_138 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_137); lean_dec(x_137); if (x_138 == 0) { @@ -75041,7 +75044,7 @@ lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; x_144 = lean_ctor_get(x_141, 0); x_145 = lean_ctor_get(x_141, 4); lean_dec(x_145); -x_146 = l_Lean_Kernel_enableDiag(x_144, x_133); +x_146 = l_Lean_Kernel_Environment_enableDiag(x_144, x_133); x_147 = l_Lean_Elab_Term_TermElabM_toIO___rarg___closed__3; lean_ctor_set(x_141, 4, x_147); lean_ctor_set(x_141, 0, x_146); @@ -75156,7 +75159,7 @@ lean_inc(x_172); lean_inc(x_171); lean_inc(x_170); lean_dec(x_141); -x_177 = l_Lean_Kernel_enableDiag(x_170, x_133); +x_177 = l_Lean_Kernel_Environment_enableDiag(x_170, x_133); x_178 = l_Lean_Elab_Term_TermElabM_toIO___rarg___closed__3; x_179 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_179, 0, x_177); @@ -75390,7 +75393,7 @@ lean_dec(x_239); x_242 = lean_ctor_get(x_240, 0); lean_inc(x_242); lean_dec(x_240); -x_243 = l_Lean_Kernel_isDiagnosticsEnabled(x_242); +x_243 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_242); lean_dec(x_242); if (x_243 == 0) { @@ -75465,7 +75468,7 @@ if (lean_is_exclusive(x_246)) { lean_dec_ref(x_246); x_255 = lean_box(0); } -x_256 = l_Lean_Kernel_enableDiag(x_248, x_238); +x_256 = l_Lean_Kernel_Environment_enableDiag(x_248, x_238); x_257 = l_Lean_Elab_Term_TermElabM_toIO___rarg___closed__3; if (lean_is_scalar(x_255)) { x_258 = lean_alloc_ctor(0, 8, 0); @@ -76944,7 +76947,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term { lean_object* x_11; uint8_t x_12; lean_inc(x_1); -lean_inc(x_2); x_11 = l_Lean_mkPrivateName(x_2, x_1); lean_inc(x_2); x_12 = l_Lean_Environment_contains(x_2, x_11); @@ -77920,6 +77922,7 @@ x_87 = lean_ctor_get(x_85, 0); lean_inc(x_87); lean_dec(x_85); x_88 = l_Lean_mkPrivateName(x_87, x_2); +lean_dec(x_87); lean_inc(x_88); x_89 = l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6(x_88, x_3, x_4, x_5, x_6, x_7, x_8, x_86); if (lean_obj_tag(x_89) == 0) @@ -80056,7 +80059,8 @@ x_16 = lean_ctor_get(x_13, 1); x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -x_18 = lean_environment_main_module(x_17); +x_18 = l_Lean_Environment_mainModule(x_17); +lean_dec(x_17); x_19 = l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__1; lean_inc(x_11); lean_ctor_set_tag(x_13, 2); @@ -80210,7 +80214,8 @@ lean_dec(x_13); x_51 = lean_ctor_get(x_49, 0); lean_inc(x_51); lean_dec(x_49); -x_52 = lean_environment_main_module(x_51); +x_52 = l_Lean_Environment_mainModule(x_51); +lean_dec(x_51); x_53 = l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__1; lean_inc(x_11); x_54 = lean_alloc_ctor(2, 2, 0); diff --git a/stage0/stdlib/Lean/Elab/Util.c b/stage0/stdlib/Lean/Elab/Util.c index 578ac6183f15..8f616ea07894 100644 --- a/stage0/stdlib/Lean/Elab/Util.c +++ b/stage0/stdlib/Lean/Elab/Util.c @@ -225,7 +225,7 @@ static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__1; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2559____closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___at_Lean_Elab_checkSyntaxNodeKindAtCurrentNamespaces___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -3643,7 +3643,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__7(lean_object* _start: { lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_environment_main_module(x_1); +x_18 = l_Lean_Environment_mainModule(x_1); x_19 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_19, 0, x_2); lean_ctor_set(x_19, 1, x_18); @@ -4102,6 +4102,7 @@ lean_object* x_17 = _args[16]; { lean_object* x_18; x_18 = l_Lean_Elab_liftMacroM___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_1); return x_18; } } diff --git a/stage0/stdlib/Lean/Environment.c b/stage0/stdlib/Lean/Environment.c index 4270ccdfa671..32abb624cf2d 100644 --- a/stage0/stdlib/Lean/Environment.c +++ b/stage0/stdlib/Lean/Environment.c @@ -13,20 +13,15 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__9___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_compacted_region_is_memory_mapped(size_t); static lean_object* l_Lean_TagDeclarationExtension_instInhabited___closed__1; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_instMonadEnvOfMonadLift(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600_(lean_object*); +static lean_object* l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__28; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6727_(lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited(lean_object*, lean_object*); @@ -35,34 +30,38 @@ static lean_object* l_Lean_MapDeclarationExtension_insert___rarg___closed__2; static lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_find_x3f___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEqOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__1; lean_object* lean_read_module_data(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_diagExt; LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Kernel_enableDiag___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkModuleData(lean_object*, lean_object*); uint32_t lean_string_utf8_get(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*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__2; +LEAN_EXPORT lean_object* lean_elab_environment_of_kernel_env(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__28; lean_object* l_Lean_Name_quickLt___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__8; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__5(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Environment_header___boxed(lean_object*); lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(lean_object*, lean_object*, lean_object*); 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_Lean_PersistentEnvExtension_modifyState___rarg___lambda__1(lean_object*, lean_object*); -lean_object* lean_add_decl(lean_object*, size_t, lean_object*, lean_object*); +lean_object* lean_elab_add_decl(lean_object*, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_Diagnostics_isEnabled___boxed(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__15; LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries(lean_object*, lean_object*); lean_object* lean_uint32_to_nat(uint32_t); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__3; +LEAN_EXPORT lean_object* lean_kernel_set_diag(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_constants(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; @@ -70,6 +69,7 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedModuleData; +LEAN_EXPORT lean_object* l_Lean_Environment_constants___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_numBuckets___at_Lean_Environment_displayStats___spec__5___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_finalizeImport(lean_object*, lean_object*, lean_object*, uint32_t, uint8_t, lean_object*); @@ -82,6 +82,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_display LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_header(lean_object*); static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); @@ -92,28 +93,27 @@ 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*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__11; +lean_object* lean_add_decl(lean_object*, size_t, lean_object*, 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*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__1; static lean_object* l_Lean_instModuleIdxBEq___closed__2; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionState; LEAN_EXPORT uint8_t l_Lean_Environment_hasUnsafe___lambda__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType___rarg___closed__1; -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__5(lean_object*, lean_object*, lean_object*); lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__19; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__3; LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_kernel_is_def_eq(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__21; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1(lean_object*, 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); @@ -135,49 +135,54 @@ uint8_t l_Lean_NameHashSet_contains(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg___closed__1; static lean_object* l_Lean_mkTagDeclarationExtension___closed__3; LEAN_EXPORT lean_object* l_Lean_readModuleData___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__26; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2; LEAN_EXPORT uint8_t l_Lean_Environment_isConstructor(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___boxed(lean_object*, lean_object*, lean_object*); -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*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3062_; -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3542_; +LEAN_EXPORT lean_object* lean_elab_environment_to_kernel_env(lean_object*); 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___auto____x40_Lean_Environment___hyg_3100____closed__4; static lean_object* l_Lean_importModules___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_name(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__2; -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* lean_environment_find(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_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_getDiagnostics___boxed(lean_object*); lean_object* lean_get_num_attributes(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023_(lean_object*); 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*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__22; -LEAN_EXPORT lean_object* lean_environment_find(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__1; lean_object* l_Std_DHashMap_Raw_Internal_numBuckets___rarg(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*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__27; LEAN_EXPORT lean_object* l_Lean_finalizeImport_unsafe__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_find_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_imports___boxed(lean_object*); static lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2___boxed(lean_object*, lean_object*, lean_object*); 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___auto____x40_Lean_Environment___hyg_2684____closed__5; 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*); -LEAN_EXPORT lean_object* lean_kernel_get_diag(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_getDiagnostics(lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportStateM_run(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_finalizeImport_unsafe__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__8(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_quickLt(lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); lean_object* lean_kernel_check(lean_object*, lean_object*, lean_object*); @@ -189,74 +194,75 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_mkInitialExtensionStates(lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_addExtraName(lean_object*, lean_object*); 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 uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionEntrySpec; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_find_x3f___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(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*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__26; 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_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4(lean_object*, lean_object*, lean_object*, 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*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_value_x21(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedImport___closed__1; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_mkExtNameMap___spec__4(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag___lambda__1(uint8_t, lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__2; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_registerNamespace(lean_object*, lean_object*); uint8_t l_Std_Format_isNil(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Kernel_Environment_isQuotInit___boxed(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__15; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__2(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* lean_elab_environment_update_base_after_kernel_add(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3(lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84_(lean_object*, lean_object*); lean_object* l_Lean_initializing(lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__22; LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* lean_environment_mark_quot_init(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_TagDeclarationExtension_tag___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__3(lean_object*, lean_object*, lean_object*, 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 uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerEnvExtension(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__7; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__9___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); 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*); static lean_object* l_Lean_instToStringImport___closed__3; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__27; -LEAN_EXPORT lean_object* l_Lean_Kernel_Diagnostics_isEnabled___boxed(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__17; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__23; lean_object* l_panic___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__2; static lean_object* l_Lean_Environment_displayStats___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___lambda__1___closed__2; static lean_object* l_Lean_Kernel_instInhabitedDiagnostics___closed__1; +LEAN_EXPORT lean_object* l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_instCoeNameImport(lean_object*); @@ -264,8 +270,10 @@ lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToStringImport(lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); @@ -284,6 +292,7 @@ static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__20; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType___rarg(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*); @@ -293,9 +302,9 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___s static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_instModuleIdxBEq; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_finalizeImport___spec__4(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_isDiagnosticsEnabled___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1(uint32_t, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___closed__1; static lean_object* l_Lean_MapDeclarationExtension_insert___rarg___closed__4; @@ -304,37 +313,40 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_fi LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Diagnostics_recordUnfold___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__25; +static lean_object* l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__5___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_resetDiag(lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__16; LEAN_EXPORT lean_object* l_Lean_instModuleIdxToString; +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__5; 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_mkEmptyEnvironment___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_addAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadEnvOfMonadLift___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_TagDeclarationExtension_tag___closed__5; -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_EnvExtensionInterfaceUnsafe_getState(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__5; LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_isDefEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_display_stats(lean_object*, lean_object*); uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__8; -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3379_; -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*); -static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__6; +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__23; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1; -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___closed__1; -static size_t l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_throwAlreadyImported___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__9; lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3958_; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2(lean_object*); @@ -342,14 +354,14 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__13; static lean_object* l_Lean_Environment_displayStats___closed__4; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__18; LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2639_(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_importModules___lambda__2(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*, lean_object*); uint8_t l_Lean_SMap_contains___at_Lean_NameSSet_contains___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__5; lean_object* lean_st_mk_ref(lean_object*, lean_object*); extern lean_object* l_Lean_NameSSet_instInhabited; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__4; @@ -361,8 +373,11 @@ static lean_object* l_Lean_instInhabitedModuleData___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_numBuckets___at_Lean_Environment_displayStats___spec__5(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_find_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg(lean_object*, lean_object*, uint32_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); @@ -370,29 +385,30 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState(lean_ob 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*); uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_431____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__13; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__3; +LEAN_EXPORT uint8_t lean_environment_quot_init(lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_addDeclCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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___auto____x40_Lean_Environment___hyg_3100____closed__7; LEAN_EXPORT lean_object* l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___closed__3; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState(lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* lean_environment_set_main_module(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_setMainModule(lean_object*, lean_object*); static lean_object* l_Lean_TagDeclarationExtension_tag___closed__3; static lean_object* l_Lean_instToStringImport___closed__1; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkDefinitionValInferrringUnsafe(lean_object*); -LEAN_EXPORT uint32_t lean_environment_trust_level(lean_object*); lean_object* l_IO_print___at_IO_println___spec__1(lean_object*, lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__2; 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_initFn____x40_Lean_Environment___hyg_6600____lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_Declaration_getNames(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -404,36 +420,32 @@ LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finaliz LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___lambda__1(lean_object*, lean_object*); lean_object* l_instToStringNat(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__6; LEAN_EXPORT lean_object* l_Lean_instToStringImport___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__2; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_instToStringImport___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__3(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__22; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__12; LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___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_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* lean_kernel_set_diag(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_setDiagnostics(lean_object*, lean_object*); lean_object* lean_update_env_attributes(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__3; static lean_object* l_Lean_TagDeclarationExtension_tag___closed__4; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__13; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___closed__4; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__23; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_isNamespaceName___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__25; LEAN_EXPORT lean_object* l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_instToStringImport___closed__2; +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__6; LEAN_EXPORT uint8_t l___private_Lean_Environment_0__Lean_Environment_isNamespaceName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_instantiateLevelParams(lean_object*, lean_object*, lean_object*); @@ -442,74 +454,78 @@ LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___la 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___auto____x40_Lean_Environment___hyg_3795_; LEAN_EXPORT lean_object* l_Lean_instInhabitedImport; 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; +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(lean_object*, lean_object*); extern lean_object* l_Std_Format_defWidth; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Environment_hasUnsafe(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_enableDiag___boxed(lean_object*, lean_object*); static lean_object* l_Lean_MapDeclarationExtension_instInhabited___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___boxed(lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__6; LEAN_EXPORT lean_object* l_Lean_registerEnvExtension___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__26; +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__24; lean_object* lean_usize_to_nat(size_t); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3100_; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_evalConstCheck___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_isDefEqGuarded___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Environment_isNamespace(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3(lean_object*, size_t, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__8; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__9(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_addDeclCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6727____closed__1; +static lean_object* l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_withImporting___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType(lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState(lean_object*); -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1472_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1(lean_object*); static lean_object* l_Lean_Environment_registerNamespace___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__17; -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__5___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed(lean_object*); +static lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1; static lean_object* l_Lean_mkTagDeclarationExtension___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionStateSpec___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* lean_environment_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* lean_elab_environment_add(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__20; LEAN_EXPORT uint8_t l_Lean_Environment_getModuleIdx_x3f___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_instInhabitedDiagnostics; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__24; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_hasUnsafe___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__1; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Environment_hasUnsafe___boxed(lean_object*, lean_object*); @@ -520,12 +536,12 @@ static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___boxed(lean_object*, lean_object*); -LEAN_EXPORT uint8_t lean_kernel_diag_is_enabled(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5___boxed(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_7023____spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); extern lean_object* l_Lean_NameSet_empty; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__10; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* lean_environment_mark_quot_init(lean_object*); lean_object* lean_string_length(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; @@ -536,22 +552,25 @@ LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg___lambda__1(le uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instReprImport; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*); 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_initFn____x40_Lean_Environment___hyg_6600____spec__2___boxed(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_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__12; LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_mainModule___boxed(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__14; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState(lean_object*, lean_object*); -LEAN_EXPORT lean_object* lean_environment_main_module(lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__4(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1833_(lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5___boxed(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__3; +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__11; LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState(lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4(size_t, 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_7023____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; @@ -561,45 +580,46 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedModuleIdx; LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_allImportedModuleNames(lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_addDeclWithoutChecking___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*); 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_object* l_Lean_findOLean(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_isConstructor___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1(lean_object*, lean_object*); -lean_object* lean_add_decl_without_checking(lean_object*, lean_object*); +lean_object* lean_elab_add_decl_without_checking(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Environment_isSafeDefinition(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__5; uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_mkTagDeclarationExtension___closed__1; LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_const2ModIdx(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); 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*); uint64_t lean_uint64_xor(uint64_t, uint64_t); extern lean_object* l_Id_instMonad; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_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*); static lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__7; static lean_object* l_Lean_throwAlreadyImported___rarg___closed__3; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__2; -static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2(lean_object*, lean_object*); uint32_t lean_uint32_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface; LEAN_EXPORT lean_object* l_Lean_mkDefinitionValInferrringUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__5; @@ -607,41 +627,41 @@ lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Kernel_isDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*); static uint32_t l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__6; LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__2(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3478_; lean_object* l_Lean_NameHashSet_insert(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry(lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_find_x3f___spec__2(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkModuleData___closed__1; static lean_object* l_Lean_TagDeclarationExtension_isTagged___closed__1; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__5; +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__19; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__3; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Kernel_resetDiag___lambda__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionStateSpec; LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_modifyState___spec__1(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_CompactedRegion_free___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__16; +LEAN_EXPORT uint8_t lean_kernel_diag_is_enabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_whnf___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_add_decl_without_checking(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__15; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__7; LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_2684_; LEAN_EXPORT lean_object* l_Lean_persistentEnvExtensionsRef; lean_object* lean_io_initializing(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__9; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__18; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__10___boxed(lean_object*, lean_object*, lean_object*, 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_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -652,13 +672,10 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_display LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Diagnostics_recordUnfold___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_namespacesExt; LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState(lean_object*); static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__1; LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_isTagged___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__4; -static size_t l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__1; static lean_object* l_Lean_mkMapDeclarationExtension___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_importModulesCore___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ModuleIdx_toNat(lean_object*); @@ -676,7 +693,6 @@ static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda_ LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedName; LEAN_EXPORT lean_object* l_Lean_Kernel_isDiagnosticsEnabled___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_isQuotInit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); @@ -684,51 +700,52 @@ 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*); 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*); LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_const2ModIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_numBuckets___at_Lean_Environment_displayStats___spec__4___boxed(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__10; LEAN_EXPORT lean_object* lean_write_module(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MapDeclarationExtension_insert___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkModuleData___spec__8(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkDefinitionValInferrringUnsafe___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1___boxed(lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t lean_environment_quot_init(lean_object*); +static lean_object* l_Lean_Kernel_instInhabitedDiagnostics___closed__3; +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__3; static lean_object* l_Lean_mkModuleData___closed__2; lean_object* l_instDecidableEqNat___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_getNamespaceSet___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__6; static lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__10; 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_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__19; static lean_object* l_Lean_Environment_displayStats___closed__2; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__3; LEAN_EXPORT lean_object* l_Lean_saveModuleData___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadEnvOfMonadLift___rarg(lean_object*, lean_object*); lean_object* lean_find_expr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* l_Lean_SMap_insert___at_Lean_NameSSet_insert___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* lean_kernel_record_unfold(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__8(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__7; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__1; uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__14; lean_object* l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_empty_environment(uint32_t, lean_object*); @@ -746,47 +763,48 @@ static lean_object* l_Lean_instModuleIdxBEq___closed__1; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState(lean_object*, lean_object*); 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___private_Lean_Environment_0__Lean_equivInfo___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* lean_environment_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_import_modules(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt(lean_object*); uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___boxed__const__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__4; LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg___boxed(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_6600____spec__8(lean_object*, size_t, size_t, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__21; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2(lean_object*); lean_object* lean_kernel_whnf(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2(lean_object*); static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_finalizeImport___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_3100____closed__16; +LEAN_EXPORT lean_object* lean_kernel_get_diag(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_TagDeclarationExtension_tag___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1(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(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_3055_(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___lambda__1___closed__3; -LEAN_EXPORT lean_object* lean_kernel_record_unfold(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__20; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains(lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__1(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); 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_Lean_Kernel_Environment_addDeclWithoutChecking___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Environment_displayStats___spec__2(lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Environment_displayStats___spec__2___closed__1; -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__6; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__4; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2684____closed__4; LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState(lean_object*); LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); @@ -795,15 +813,18 @@ LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1___boxed(lean_obje LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__6(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_84____closed__14; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__9(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(lean_object*, lean_object*, lean_object*, 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; LEAN_EXPORT lean_object* l_Lean_ModuleIdx_toNat___boxed(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedModuleData___closed__2; LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported(lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__4(size_t, 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_7023____spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* _init_l_Lean_EnvExtensionStateSpec___closed__1() { _start: { @@ -1420,134 +1441,87 @@ x_1 = l_Lean_instInhabitedModuleData___closed__2; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_Kernel_instInhabitedDiagnostics___closed__1() { _start: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_5); -return x_6; +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +return x_1; } -else -{ -lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Name_hash___override(x_9); -x_12 = lean_uint64_to_usize(x_11); -x_13 = 1; -x_14 = lean_usize_sub(x_1, x_13); -x_15 = 5; -x_16 = lean_usize_mul(x_15, x_14); -x_17 = lean_usize_shift_right(x_12, x_16); -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_5, x_18); -lean_dec(x_5); -x_20 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(x_6, x_17, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_19; -x_6 = x_20; -goto _start; } +static lean_object* _init_l_Lean_Kernel_instInhabitedDiagnostics___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Kernel_instInhabitedDiagnostics___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Kernel_instInhabitedDiagnostics___closed__3() { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_1, 1); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 0); -lean_dec(x_11); -x_12 = lean_array_push(x_5, x_3); -x_13 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_12); -return x_1; +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +x_2 = 0; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Kernel_instInhabitedDiagnostics() { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_1); -x_14 = lean_array_push(x_5, x_3); -x_15 = lean_array_push(x_6, x_4); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_object* x_1; +x_1 = l_Lean_Kernel_instInhabitedDiagnostics___closed__3; +return x_1; } } -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_name_eq(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_find_x3f___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_19; lean_object* x_20; +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); lean_dec(x_6); -lean_dec(x_5); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_2, x_19); -lean_dec(x_2); -x_2 = x_20; -goto _start; +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -if (x_22 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_name_eq(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 1); -lean_dec(x_23); -x_24 = lean_ctor_get(x_1, 0); -lean_dec(x_24); -x_25 = lean_array_fset(x_5, x_2, x_3); -x_26 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -lean_ctor_set(x_1, 1, x_26); -lean_ctor_set(x_1, 0, x_25); -return x_1; +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_1); -x_27 = lean_array_fset(x_5, x_2, x_3); -x_28 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} +lean_object* x_14; lean_object* x_15; +x_14 = lean_array_fget(x_2, x_4); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; } } } } -static size_t _init_l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__1() { +static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__1() { _start: { size_t x_1; size_t x_2; size_t x_3; @@ -1557,1872 +1531,1688 @@ x_3 = lean_usize_shift_left(x_1, x_2); return x_3; } } -static size_t _init_l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2() { +static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2() { _start: { size_t x_1; size_t x_2; size_t x_3; x_1 = 1; -x_2 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__1; +x_2 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__1; x_3 = lean_usize_sub(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) { -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = 1; -x_9 = 5; -x_10 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_14 = lean_nat_dec_lt(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) { -lean_dec(x_12); +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); lean_dec(x_5); -lean_dec(x_4); -return x_1; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { +switch (lean_obj_tag(x_11)) { case 0: { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_name_eq(x_4, x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_free_object(x_15); -x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_array_fset(x_17, x_12, x_23); +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_name_eq(x_3, x_12); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_24); -return x_1; +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; } else { -lean_object* x_25; -lean_dec(x_20); -lean_dec(x_19); -lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_4); -x_25 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_25); +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); return x_1; } } -else -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_name_eq(x_4, x_26); -if (x_28 == 0) +case 1: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_array_fset(x_17, x_12, x_30); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); -return x_1; +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; +goto _start; } -else +default: { -lean_object* x_32; lean_object* x_33; -lean_dec(x_27); -lean_dec(x_26); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_4); -lean_ctor_set(x_32, 1, x_5); -x_33 = lean_array_fset(x_17, x_12, x_32); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_33); -return x_1; +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); +return x_19; } } } -case 1: +else { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { +case 0: { -lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -x_35 = lean_ctor_get(x_15, 0); -x_36 = lean_usize_shift_right(x_2, x_9); -x_37 = lean_usize_add(x_3, x_8); -x_38 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(x_35, x_36, x_37, x_4, x_5); -lean_ctor_set(x_15, 0, x_38); -x_39 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); -return x_1; +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_name_eq(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; } else { -lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_15, 0); -lean_inc(x_40); -lean_dec(x_15); -x_41 = lean_usize_shift_right(x_2, x_9); -x_42 = lean_usize_add(x_3, x_8); -x_43 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(x_40, x_41, x_42, x_4, x_5); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_array_fset(x_17, x_12, x_44); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); -return x_1; +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; } } +case 1: +{ +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; +} default: { -lean_object* x_46; lean_object* x_47; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_4); -lean_ctor_set(x_46, 1, x_5); -x_47 = lean_array_fset(x_17, x_12, x_46); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_47); -return x_1; +lean_object* x_35; +x_35 = lean_box(0); +return x_35; } } } } else { -lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_1, 0); -lean_inc(x_48); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); lean_dec(x_1); -x_49 = 1; -x_50 = 5; -x_51 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_52 = lean_usize_land(x_2, x_51); -x_53 = lean_usize_to_nat(x_52); -x_54 = lean_array_get_size(x_48); -x_55 = lean_nat_dec_lt(x_53, x_54); -lean_dec(x_54); -if (x_55 == 0) +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_find_x3f___spec__4(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_find_x3f___spec__2(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_56; -lean_dec(x_53); -lean_dec(x_5); -lean_dec(x_4); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_48); -return x_56; +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3(x_1, x_4, x_2); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_array_fget(x_48, x_53); -x_58 = lean_box(0); -x_59 = lean_array_fset(x_48, x_53, x_58); -switch (lean_obj_tag(x_57)) { -case 0: +if (lean_obj_tag(x_2) == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_57, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_62 = x_57; -} else { - lean_dec_ref(x_57); - x_62 = lean_box(0); +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } -x_63 = lean_name_eq(x_4, x_60); -if (x_63 == 0) +else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_array_fset(x_59, x_53, x_65); -lean_dec(x_53); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_name_eq(x_4, x_1); +if (x_7 == 0) +{ +x_2 = x_6; +goto _start; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_61); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_68 = lean_alloc_ctor(0, 2, 0); -} else { - x_68 = x_62; -} -lean_ctor_set(x_68, 0, x_4); -lean_ctor_set(x_68, 1, x_5); -x_69 = lean_array_fset(x_59, x_53, x_68); -lean_dec(x_53); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; -} +lean_object* x_9; +lean_inc(x_5); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_5); +return x_9; } -case 1: -{ -lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_71 = lean_ctor_get(x_57, 0); -lean_inc(x_71); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - x_72 = x_57; -} else { - lean_dec_ref(x_57); - x_72 = lean_box(0); } -x_73 = lean_usize_shift_right(x_2, x_50); -x_74 = lean_usize_add(x_3, x_49); -x_75 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(x_71, x_73, x_74, x_4, x_5); -if (lean_is_scalar(x_72)) { - x_76 = lean_alloc_ctor(1, 1, 0); -} else { - x_76 = x_72; } -lean_ctor_set(x_76, 0, x_75); -x_77 = lean_array_fset(x_59, x_53, x_76); -lean_dec(x_53); -x_78 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_78, 0, x_77); -return x_78; } -default: +LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_4); -lean_ctor_set(x_79, 1, x_5); -x_80 = lean_array_fset(x_59, x_53, x_79); -lean_dec(x_53); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; -} -} -} -} -} -else +uint8_t x_3; +x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +if (x_3 == 0) { -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_array_get_size(x_6); +x_8 = l_Lean_Name_hash___override(x_2); +x_9 = 32; +x_10 = lean_uint64_shift_right(x_8, x_9); +x_11 = lean_uint64_xor(x_8, x_10); +x_12 = 16; +x_13 = lean_uint64_shift_right(x_11, x_12); +x_14 = lean_uint64_xor(x_11, x_13); +x_15 = lean_uint64_to_usize(x_14); +x_16 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_17 = 1; +x_18 = lean_usize_sub(x_16, x_17); +x_19 = lean_usize_land(x_15, x_18); +x_20 = lean_array_uget(x_6, x_19); +lean_dec(x_6); +x_21 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_2, x_20); +lean_dec(x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; -x_83 = lean_unsigned_to_nat(0u); -x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__5(x_1, x_83, x_4, x_5); -x_85 = 7; -x_86 = lean_usize_dec_le(x_85, x_3); -if (x_86 == 0) +lean_object* x_22; +x_22 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_find_x3f___spec__2(x_5, x_2); +return x_22; +} +else { -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_dec_lt(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) +uint8_t x_23; +lean_dec(x_5); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_84, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 1); -lean_inc(x_91); -lean_dec(x_84); -x_92 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4(x_3, x_90, x_91, lean_box(0), x_83, x_92); -lean_dec(x_91); -lean_dec(x_90); -return x_93; +return x_21; } else { -return x_84; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +return x_25; +} } } else { -return x_84; +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_array_get_size(x_27); +x_29 = l_Lean_Name_hash___override(x_2); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_28); +lean_dec(x_28); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_27, x_40); +lean_dec(x_27); +x_42 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_2, x_41); +lean_dec(x_41); +return x_42; } } -else +} +LEAN_EXPORT lean_object* lean_environment_find(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; -x_94 = lean_ctor_get(x_1, 0); -x_95 = lean_ctor_get(x_1, 1); -lean_inc(x_95); -lean_inc(x_94); +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); lean_dec(x_1); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -x_97 = lean_unsigned_to_nat(0u); -x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__5(x_96, x_97, x_4, x_5); -x_99 = 7; -x_100 = lean_usize_dec_le(x_99, x_3); -if (x_100 == 0) -{ -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); -x_102 = lean_unsigned_to_nat(4u); -x_103 = lean_nat_dec_lt(x_101, x_102); -lean_dec(x_101); -if (x_103 == 0) -{ -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_98, 1); -lean_inc(x_105); -lean_dec(x_98); -x_106 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; +x_4 = l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(x_3, x_2); +lean_dec(x_2); +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_find_x3f___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -return x_98; +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_find_x3f___spec__4(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_6; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -return x_98; +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; +} } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_find_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_find_x3f___spec__2(x_1, x_2); +lean_dec(x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Name_hash___override(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(x_1, x_5, x_6, x_2, x_3); -return x_7; +lean_object* x_3; +x_3 = l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; } } -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* lean_environment_mark_quot_init(lean_object* x_1) { _start: { -if (lean_obj_tag(x_2) == 0) +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) { uint8_t x_3; -x_3 = 0; -return x_3; +x_3 = 1; +lean_ctor_set_uint8(x_1, sizeof(void*)*6, x_3); +return x_1; } else { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_name_eq(x_4, x_1); -if (x_6 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_ctor_get(x_1, 3); +x_8 = lean_ctor_get(x_1, 4); +x_9 = lean_ctor_get(x_1, 5); +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_dec(x_1); +x_10 = 1; +x_11 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_5); +lean_ctor_set(x_11, 2, x_6); +lean_ctor_set(x_11, 3, x_7); +lean_ctor_set(x_11, 4, x_8); +lean_ctor_set(x_11, 5, x_9); +lean_ctor_set_uint8(x_11, sizeof(void*)*6, x_10); +return x_11; +} +} +} +LEAN_EXPORT uint8_t lean_environment_quot_init(lean_object* x_1) { +_start: { -x_2 = x_5; -goto _start; +uint8_t x_2; +x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +lean_dec(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Kernel_Environment_isQuotInit___boxed(lean_object* x_1) { +_start: { -uint8_t x_8; -x_8 = 1; -return x_8; +uint8_t x_2; lean_object* x_3; +x_2 = lean_environment_quot_init(x_1); +x_3 = lean_box(x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_addDeclCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; lean_object* x_6; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_add_decl(x_1, x_5, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_6; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__9(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_addDeclWithoutChecking___boxed(lean_object* x_1, lean_object* x_2) { _start: { -if (lean_obj_tag(x_2) == 0) +lean_object* x_3; +x_3 = lean_add_decl_without_checking(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -return x_1; +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_dec(x_5); +return x_6; } else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = l_Lean_Name_hash___override(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; +lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Name_hash___override(x_9); +x_12 = lean_uint64_to_usize(x_11); +x_13 = 1; +x_14 = lean_usize_sub(x_1, x_13); +x_15 = 5; +x_16 = lean_usize_mul(x_15, x_14); +x_17 = lean_usize_shift_right(x_12, x_16); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_5, x_18); +lean_dec(x_5); +x_20 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3(x_6, x_17, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_19; +x_6 = x_20; goto _start; } -else +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = l_Lean_Name_hash___override(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; } +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_name_eq(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_6); +lean_dec(x_5); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_2); +x_2 = x_20; +goto _start; } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 0) { +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_1, 0); +lean_dec(x_24); +x_25 = lean_array_fset(x_5, x_2, x_3); +x_26 = lean_array_fset(x_6, x_2, x_4); lean_dec(x_2); -lean_dec(x_1); -return x_3; +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); +return x_1; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__9(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(lean_object* x_1) { +static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1() { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__8(x_7, x_1, x_6); -return x_8; +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -if (lean_obj_tag(x_3) == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_4; -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; } else { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 1); -x_8 = lean_ctor_get(x_3, 2); -x_9 = lean_name_eq(x_6, x_1); -if (x_9 == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_10; -x_10 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_10); -return x_3; +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_name_eq(x_4, x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_array_fset(x_17, x_12, x_23); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_24); +return x_1; } else { -lean_dec(x_7); -lean_dec(x_6); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_1); -return x_3; +lean_object* x_25; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_25 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_25); +return x_1; } } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_3, 0); -x_12 = lean_ctor_get(x_3, 1); -x_13 = lean_ctor_get(x_3, 2); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_3); -x_14 = lean_name_eq(x_11, x_1); -if (x_14 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_name_eq(x_4, x_26); +if (x_28 == 0) { -lean_object* x_15; lean_object* x_16; -x_15 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_1, x_2, x_13); -x_16 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_16, 0, x_11); -lean_ctor_set(x_16, 1, x_12); -lean_ctor_set(x_16, 2, x_15); -return x_16; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = lean_array_fset(x_17, x_12, x_30); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_31); +return x_1; } else { -lean_object* x_17; +lean_object* x_32; lean_object* x_33; +lean_dec(x_27); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_5); +x_33 = lean_array_fset(x_17, x_12, x_32); lean_dec(x_12); -lean_dec(x_11); -x_17 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_2); -lean_ctor_set(x_17, 2, x_13); -return x_17; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -if (x_4 == 0) -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_1); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_1, 1); -x_7 = l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__2(x_6, x_2, x_3); -x_8 = 0; -lean_ctor_set(x_1, 1, x_7); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_8); +lean_ctor_set(x_1, 0, x_33); return x_1; } -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__2(x_10, x_2, x_3); -x_12 = 0; -x_13 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_13, 0, x_9); -lean_ctor_set(x_13, 1, x_11); -lean_ctor_set_uint8(x_13, sizeof(void*)*2, x_12); -return x_13; } } -else -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_1); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_1, 0); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; lean_object* x_32; uint8_t x_33; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -x_19 = lean_array_get_size(x_18); -x_20 = l_Lean_Name_hash___override(x_2); -x_21 = 32; -x_22 = lean_uint64_shift_right(x_20, x_21); -x_23 = lean_uint64_xor(x_20, x_22); -x_24 = 16; -x_25 = lean_uint64_shift_right(x_23, x_24); -x_26 = lean_uint64_xor(x_23, x_25); -x_27 = lean_uint64_to_usize(x_26); -x_28 = lean_usize_of_nat(x_19); -lean_dec(x_19); -x_29 = 1; -x_30 = lean_usize_sub(x_28, x_29); -x_31 = lean_usize_land(x_27, x_30); -x_32 = lean_array_uget(x_18, x_31); -x_33 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_2, x_32); -if (x_33 == 0) +case 1: { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_add(x_17, x_34); -lean_dec(x_17); -x_36 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_36, 0, x_2); -lean_ctor_set(x_36, 1, x_3); -lean_ctor_set(x_36, 2, x_32); -x_37 = lean_array_uset(x_18, x_31, x_36); -x_38 = lean_unsigned_to_nat(4u); -x_39 = lean_nat_mul(x_35, x_38); -x_40 = lean_unsigned_to_nat(3u); -x_41 = lean_nat_div(x_39, x_40); -lean_dec(x_39); -x_42 = lean_array_get_size(x_37); -x_43 = lean_nat_dec_le(x_41, x_42); -lean_dec(x_42); -lean_dec(x_41); -if (x_43 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -lean_object* x_44; uint8_t x_45; -x_44 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_37); -lean_ctor_set(x_15, 1, x_44); -lean_ctor_set(x_15, 0, x_35); -x_45 = 1; -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_45); +lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_usize_shift_right(x_2, x_9); +x_37 = lean_usize_add(x_3, x_8); +x_38 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3(x_35, x_36, x_37, x_4, x_5); +lean_ctor_set(x_15, 0, x_38); +x_39 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); return x_1; } else { -uint8_t x_46; -lean_ctor_set(x_15, 1, x_37); -lean_ctor_set(x_15, 0, x_35); -x_46 = 1; -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_46); +lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_15, 0); +lean_inc(x_40); +lean_dec(x_15); +x_41 = lean_usize_shift_right(x_2, x_9); +x_42 = lean_usize_add(x_3, x_8); +x_43 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3(x_40, x_41, x_42, x_4, x_5); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_array_fset(x_17, x_12, x_44); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_45); return x_1; } } -else +default: { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_47 = lean_box(0); -x_48 = lean_array_uset(x_18, x_31, x_47); -x_49 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_2, x_3, x_32); -x_50 = lean_array_uset(x_48, x_31, x_49); -lean_ctor_set(x_15, 1, x_50); -x_51 = 1; -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_51); +lean_object* x_46; lean_object* x_47; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_4); +lean_ctor_set(x_46, 1, x_5); +x_47 = lean_array_fset(x_17, x_12, x_46); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_47); return x_1; } } +} +} else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; uint64_t x_61; size_t x_62; size_t x_63; size_t x_64; size_t x_65; size_t x_66; lean_object* x_67; uint8_t x_68; -x_52 = lean_ctor_get(x_15, 0); -x_53 = lean_ctor_get(x_15, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_15); -x_54 = lean_array_get_size(x_53); -x_55 = l_Lean_Name_hash___override(x_2); -x_56 = 32; -x_57 = lean_uint64_shift_right(x_55, x_56); -x_58 = lean_uint64_xor(x_55, x_57); -x_59 = 16; -x_60 = lean_uint64_shift_right(x_58, x_59); -x_61 = lean_uint64_xor(x_58, x_60); -x_62 = lean_uint64_to_usize(x_61); -x_63 = lean_usize_of_nat(x_54); +lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +lean_dec(x_1); +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_52 = lean_usize_land(x_2, x_51); +x_53 = lean_usize_to_nat(x_52); +x_54 = lean_array_get_size(x_48); +x_55 = lean_nat_dec_lt(x_53, x_54); lean_dec(x_54); -x_64 = 1; -x_65 = lean_usize_sub(x_63, x_64); -x_66 = lean_usize_land(x_62, x_65); -x_67 = lean_array_uget(x_53, x_66); -x_68 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_2, x_67); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_69 = lean_unsigned_to_nat(1u); -x_70 = lean_nat_add(x_52, x_69); -lean_dec(x_52); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_2); -lean_ctor_set(x_71, 1, x_3); -lean_ctor_set(x_71, 2, x_67); -x_72 = lean_array_uset(x_53, x_66, x_71); -x_73 = lean_unsigned_to_nat(4u); -x_74 = lean_nat_mul(x_70, x_73); -x_75 = lean_unsigned_to_nat(3u); -x_76 = lean_nat_div(x_74, x_75); -lean_dec(x_74); -x_77 = lean_array_get_size(x_72); -x_78 = lean_nat_dec_le(x_76, x_77); -lean_dec(x_77); -lean_dec(x_76); -if (x_78 == 0) +if (x_55 == 0) { -lean_object* x_79; lean_object* x_80; uint8_t x_81; -x_79 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_72); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_70); -lean_ctor_set(x_80, 1, x_79); -x_81 = 1; -lean_ctor_set(x_1, 0, x_80); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_81); -return x_1; +lean_object* x_56; +lean_dec(x_53); +lean_dec(x_5); +lean_dec(x_4); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; } else { -lean_object* x_82; uint8_t x_83; -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_70); -lean_ctor_set(x_82, 1, x_72); -x_83 = 1; -lean_ctor_set(x_1, 0, x_82); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_83); -return x_1; -} -} -else +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_array_fget(x_48, x_53); +x_58 = lean_box(0); +x_59 = lean_array_fset(x_48, x_53, x_58); +switch (lean_obj_tag(x_57)) { +case 0: { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_84 = lean_box(0); -x_85 = lean_array_uset(x_53, x_66, x_84); -x_86 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_2, x_3, x_67); -x_87 = lean_array_uset(x_85, x_66, x_86); -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_52); -lean_ctor_set(x_88, 1, x_87); -x_89 = 1; -lean_ctor_set(x_1, 0, x_88); -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_89); -return x_1; -} +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_62 = x_57; +} else { + lean_dec_ref(x_57); + x_62 = lean_box(0); } +x_63 = lean_name_eq(x_4, x_60); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_66 = lean_array_fset(x_59, x_53, x_65); +lean_dec(x_53); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint64_t x_96; uint64_t x_97; uint64_t x_98; uint64_t x_99; uint64_t x_100; uint64_t x_101; uint64_t x_102; size_t x_103; size_t x_104; size_t x_105; size_t x_106; size_t x_107; lean_object* x_108; uint8_t x_109; -x_90 = lean_ctor_get(x_1, 0); -x_91 = lean_ctor_get(x_1, 1); -lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_1); -x_92 = lean_ctor_get(x_90, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_90, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_94 = x_90; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_61); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_68 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_90); - x_94 = lean_box(0); + x_68 = x_62; } -x_95 = lean_array_get_size(x_93); -x_96 = l_Lean_Name_hash___override(x_2); -x_97 = 32; -x_98 = lean_uint64_shift_right(x_96, x_97); -x_99 = lean_uint64_xor(x_96, x_98); -x_100 = 16; -x_101 = lean_uint64_shift_right(x_99, x_100); -x_102 = lean_uint64_xor(x_99, x_101); -x_103 = lean_uint64_to_usize(x_102); -x_104 = lean_usize_of_nat(x_95); -lean_dec(x_95); -x_105 = 1; -x_106 = lean_usize_sub(x_104, x_105); -x_107 = lean_usize_land(x_103, x_106); -x_108 = lean_array_uget(x_93, x_107); -x_109 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_2, x_108); -if (x_109 == 0) -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; -x_110 = lean_unsigned_to_nat(1u); -x_111 = lean_nat_add(x_92, x_110); -lean_dec(x_92); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_2); -lean_ctor_set(x_112, 1, x_3); -lean_ctor_set(x_112, 2, x_108); -x_113 = lean_array_uset(x_93, x_107, x_112); -x_114 = lean_unsigned_to_nat(4u); -x_115 = lean_nat_mul(x_111, x_114); -x_116 = lean_unsigned_to_nat(3u); -x_117 = lean_nat_div(x_115, x_116); -lean_dec(x_115); -x_118 = lean_array_get_size(x_113); -x_119 = lean_nat_dec_le(x_117, x_118); -lean_dec(x_118); -lean_dec(x_117); -if (x_119 == 0) -{ -lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; -x_120 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_113); -if (lean_is_scalar(x_94)) { - x_121 = lean_alloc_ctor(0, 2, 0); -} else { - x_121 = x_94; +lean_ctor_set(x_68, 0, x_4); +lean_ctor_set(x_68, 1, x_5); +x_69 = lean_array_fset(x_59, x_53, x_68); +lean_dec(x_53); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; } -lean_ctor_set(x_121, 0, x_111); -lean_ctor_set(x_121, 1, x_120); -x_122 = 1; -x_123 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_91); -lean_ctor_set_uint8(x_123, sizeof(void*)*2, x_122); -return x_123; } -else +case 1: { -lean_object* x_124; uint8_t x_125; lean_object* x_126; -if (lean_is_scalar(x_94)) { - x_124 = lean_alloc_ctor(0, 2, 0); +lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_71 = lean_ctor_get(x_57, 0); +lean_inc(x_71); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_72 = x_57; } else { - x_124 = x_94; + lean_dec_ref(x_57); + x_72 = lean_box(0); } -lean_ctor_set(x_124, 0, x_111); -lean_ctor_set(x_124, 1, x_113); -x_125 = 1; -x_126 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_126, 0, x_124); -lean_ctor_set(x_126, 1, x_91); -lean_ctor_set_uint8(x_126, sizeof(void*)*2, x_125); -return x_126; +x_73 = lean_usize_shift_right(x_2, x_50); +x_74 = lean_usize_add(x_3, x_49); +x_75 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3(x_71, x_73, x_74, x_4, x_5); +if (lean_is_scalar(x_72)) { + x_76 = lean_alloc_ctor(1, 1, 0); +} else { + x_76 = x_72; } +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_array_fset(x_59, x_53, x_76); +lean_dec(x_53); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; } -else +default: { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; -x_127 = lean_box(0); -x_128 = lean_array_uset(x_93, x_107, x_127); -x_129 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_2, x_3, x_108); -x_130 = lean_array_uset(x_128, x_107, x_129); -if (lean_is_scalar(x_94)) { - x_131 = lean_alloc_ctor(0, 2, 0); -} else { - x_131 = x_94; -} -lean_ctor_set(x_131, 0, x_92); -lean_ctor_set(x_131, 1, x_130); -x_132 = 1; -x_133 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_91); -lean_ctor_set_uint8(x_133, sizeof(void*)*2, x_132); -return x_133; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_4); +lean_ctor_set(x_79, 1, x_5); +x_80 = lean_array_fset(x_59, x_53, x_79); +lean_dec(x_53); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_addAux(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_1, 1); -x_5 = l_Lean_ConstantInfo_name(x_2); -x_6 = l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__1(x_4, x_5, x_2); -lean_ctor_set(x_1, 1, x_6); -return x_1; -} -else +lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; +x_83 = lean_unsigned_to_nat(0u); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__5(x_1, x_83, x_4, x_5); +x_85 = 7; +x_86 = lean_usize_dec_le(x_85, x_3); +if (x_86 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = lean_ctor_get(x_1, 1); -x_9 = lean_ctor_get(x_1, 2); -x_10 = lean_ctor_get(x_1, 3); -x_11 = lean_ctor_get(x_1, 4); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_1); -x_12 = l_Lean_ConstantInfo_name(x_2); -x_13 = l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__1(x_8, x_12, x_2); -x_14 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_14, 0, x_7); -lean_ctor_set(x_14, 1, x_13); -lean_ctor_set(x_14, 2, x_9); -lean_ctor_set(x_14, 3, x_10); -lean_ctor_set(x_14, 4, x_11); -return x_14; -} -} +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_dec_lt(x_87, x_88); +lean_dec(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_84, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 1); +lean_inc(x_91); +lean_dec(x_84); +x_92 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__4(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; lean_object* x_8; -x_7 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__4(x_7, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_3); -lean_dec(x_2); -return x_8; +return x_84; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3(x_1, x_6, x_7, x_4, x_5); -return x_8; +return x_84; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_1, x_2); -lean_dec(x_2); +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); +lean_inc(x_94); lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_unsigned_to_nat(0u); +x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__5(x_96, x_97, x_4, x_5); +x_99 = 7; +x_100 = lean_usize_dec_le(x_99, x_3); +if (x_100 == 0) +{ +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_dec_lt(x_101, x_102); +lean_dec(x_101); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__4(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; } +else +{ +return x_98; } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +else +{ +return x_98; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); -if (x_7 == 0) +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Name_hash___override(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3(x_1, x_5, x_6, x_2, x_3); +return x_7; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_8; -lean_dec(x_4); -x_8 = 0; -return x_8; +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_name_eq(x_5, x_9); -lean_dec(x_9); -if (x_10 == 0) +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_name_eq(x_4, x_1); +if (x_6 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_12; +x_2 = x_5; goto _start; } else { -uint8_t x_14; -lean_dec(x_4); -x_14 = 1; -return x_14; +uint8_t x_8; +x_8 = 1; +return x_8; } } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__9(lean_object* x_1, lean_object* x_2) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_7 = lean_usize_land(x_2, x_6); -x_8 = lean_usize_to_nat(x_7); -x_9 = lean_box(2); -x_10 = lean_array_get(x_9, x_4, x_8); -lean_dec(x_8); -lean_dec(x_4); -switch (lean_obj_tag(x_10)) { -case 0: +if (lean_obj_tag(x_2) == 0) { -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_name_eq(x_3, x_11); -lean_dec(x_11); -return x_12; +return x_1; } -case 1: +else { -lean_object* x_13; size_t x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_usize_shift_right(x_2, x_5); -x_1 = x_13; -x_2 = x_14; +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = l_Lean_Name_hash___override(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; goto _start; } -default: +else { -uint8_t x_16; -x_16 = 0; -return x_16; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = l_Lean_Name_hash___override(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} } } } +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 1); -lean_inc(x_18); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__9(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); lean_dec(x_1); -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4(x_17, x_18, lean_box(0), x_19, x_3); -lean_dec(x_18); -lean_dec(x_17); -return x_20; +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(lean_object* x_1) { _start: { -uint64_t x_3; size_t x_4; uint8_t x_5; -x_3 = l_Lean_Name_hash___override(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3(x_1, x_4, x_2); -return x_5; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__8(x_7, x_1, x_6); +return x_8; } } -LEAN_EXPORT uint8_t l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_3; -x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -if (x_3 == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; lean_object* x_20; uint8_t x_21; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); +lean_object* x_4; +lean_dec(x_2); lean_dec(x_1); -x_6 = lean_ctor_get(x_4, 1); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_array_get_size(x_6); -x_8 = l_Lean_Name_hash___override(x_2); -x_9 = 32; -x_10 = lean_uint64_shift_right(x_8, x_9); -x_11 = lean_uint64_xor(x_8, x_10); -x_12 = 16; -x_13 = lean_uint64_shift_right(x_11, x_12); -x_14 = lean_uint64_xor(x_11, x_13); -x_15 = lean_uint64_to_usize(x_14); -x_16 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_17 = 1; -x_18 = lean_usize_sub(x_16, x_17); -x_19 = lean_usize_land(x_15, x_18); -x_20 = lean_array_uget(x_6, x_19); -lean_dec(x_6); -x_21 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_2, x_20); -lean_dec(x_20); -if (x_21 == 0) -{ -uint8_t x_22; -x_22 = l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(x_5, x_2); -return x_22; -} -else -{ -uint8_t x_23; -lean_dec(x_5); -x_23 = 1; -return x_23; -} +x_4 = lean_box(0); +return x_4; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; size_t x_38; lean_object* x_39; uint8_t x_40; -x_24 = lean_ctor_get(x_1, 0); -lean_inc(x_24); -lean_dec(x_1); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_array_get_size(x_25); -x_27 = l_Lean_Name_hash___override(x_2); -x_28 = 32; -x_29 = lean_uint64_shift_right(x_27, x_28); -x_30 = lean_uint64_xor(x_27, x_29); -x_31 = 16; -x_32 = lean_uint64_shift_right(x_30, x_31); -x_33 = lean_uint64_xor(x_30, x_32); -x_34 = lean_uint64_to_usize(x_33); -x_35 = lean_usize_of_nat(x_26); -lean_dec(x_26); -x_36 = 1; -x_37 = lean_usize_sub(x_35, x_36); -x_38 = lean_usize_land(x_34, x_37); -x_39 = lean_array_uget(x_25, x_38); -lean_dec(x_25); -x_40 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_2, x_39); -lean_dec(x_39); -return x_40; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Environment_addExtraName(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 2); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 3); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 4); -lean_inc(x_7); -lean_inc(x_4); -x_8 = l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(x_4, x_2); -if (x_8 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_1); +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_name_eq(x_6, x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_ctor_get(x_1, 4); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 3); -lean_dec(x_11); -x_12 = lean_ctor_get(x_1, 2); -lean_dec(x_12); -x_13 = lean_ctor_get(x_1, 1); -lean_dec(x_13); -x_14 = lean_ctor_get(x_1, 0); -lean_dec(x_14); -x_15 = lean_box(0); -x_16 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_6, x_2, x_15); -lean_ctor_set(x_1, 3, x_16); -return x_1; +lean_object* x_10; +x_10 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_10); +return x_3; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec(x_1); -x_17 = lean_box(0); -x_18 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_6, x_2, x_17); -x_19 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_19, 0, x_3); -lean_ctor_set(x_19, 1, x_4); -lean_ctor_set(x_19, 2, x_5); -lean_ctor_set(x_19, 3, x_18); -lean_ctor_set(x_19, 4, x_7); -return x_19; +lean_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } else { -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_3, 2); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); lean_dec(x_3); -lean_dec(x_2); -return x_1; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +x_14 = lean_name_eq(x_11, x_1); +if (x_14 == 0) { -uint8_t x_6; lean_object* x_7; -x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; -} +lean_object* x_15; lean_object* x_16; +x_15 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(x_1, x_2, x_13); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_15); +return x_16; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -size_t x_4; uint8_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3(x_1, x_4, x_3); -lean_dec(x_3); -x_6 = lean_box(x_5); -return x_6; -} +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 2, x_13); +return x_17; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(x_1, x_2); -lean_dec(x_2); -x_4 = lean_box(x_3); -return x_4; } } -LEAN_EXPORT lean_object* l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(x_1, x_2); -lean_dec(x_2); -x_4 = lean_box(x_3); -return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); -if (x_7 == 0) -{ -lean_object* x_8; -lean_dec(x_4); -x_8 = lean_box(0); -return x_8; -} -else +uint8_t x_4; +x_4 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +if (x_4 == 0) { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_name_eq(x_5, x_9); -lean_dec(x_9); -if (x_10 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_1); +if (x_5 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_12; -goto _start; +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__2(x_6, x_2, x_3); +x_8 = 0; +lean_ctor_set(x_1, 1, x_7); +lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_8); +return x_1; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_array_fget(x_2, x_4); -lean_dec(x_4); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; -} -} +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_PersistentHashMap_insert___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__2(x_10, x_2, x_3); +x_12 = 0; +x_13 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_13, 0, x_9); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*2, x_12); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_box(2); -x_11 = lean_array_get(x_10, x_5, x_9); -lean_dec(x_9); -lean_dec(x_5); -switch (lean_obj_tag(x_11)) { -case 0: +else { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_name_eq(x_3, x_12); -lean_dec(x_12); +uint8_t x_14; +x_14 = !lean_is_exclusive(x_1); if (x_14 == 0) { -lean_object* x_15; -lean_dec(x_13); -lean_free_object(x_1); -x_15 = lean_box(0); -return x_15; +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_1, 0); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; lean_object* x_32; uint8_t x_33; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = lean_array_get_size(x_18); +x_20 = l_Lean_Name_hash___override(x_2); +x_21 = 32; +x_22 = lean_uint64_shift_right(x_20, x_21); +x_23 = lean_uint64_xor(x_20, x_22); +x_24 = 16; +x_25 = lean_uint64_shift_right(x_23, x_24); +x_26 = lean_uint64_xor(x_23, x_25); +x_27 = lean_uint64_to_usize(x_26); +x_28 = lean_usize_of_nat(x_19); +lean_dec(x_19); +x_29 = 1; +x_30 = lean_usize_sub(x_28, x_29); +x_31 = lean_usize_land(x_27, x_30); +x_32 = lean_array_uget(x_18, x_31); +x_33 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_2, x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_add(x_17, x_34); +lean_dec(x_17); +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_2); +lean_ctor_set(x_36, 1, x_3); +lean_ctor_set(x_36, 2, x_32); +x_37 = lean_array_uset(x_18, x_31, x_36); +x_38 = lean_unsigned_to_nat(4u); +x_39 = lean_nat_mul(x_35, x_38); +x_40 = lean_unsigned_to_nat(3u); +x_41 = lean_nat_div(x_39, x_40); +lean_dec(x_39); +x_42 = lean_array_get_size(x_37); +x_43 = lean_nat_dec_le(x_41, x_42); +lean_dec(x_42); +lean_dec(x_41); +if (x_43 == 0) +{ +lean_object* x_44; uint8_t x_45; +x_44 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_37); +lean_ctor_set(x_15, 1, x_44); +lean_ctor_set(x_15, 0, x_35); +x_45 = 1; +lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_45); +return x_1; } else { -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_13); +uint8_t x_46; +lean_ctor_set(x_15, 1, x_37); +lean_ctor_set(x_15, 0, x_35); +x_46 = 1; +lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_46); return x_1; } } -case 1: -{ -lean_object* x_16; size_t x_17; -lean_free_object(x_1); -x_16 = lean_ctor_get(x_11, 0); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_usize_shift_right(x_2, x_6); -x_1 = x_16; -x_2 = x_17; -goto _start; -} -default: +else { -lean_object* x_19; -lean_free_object(x_1); -x_19 = lean_box(0); -return x_19; -} +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_47 = lean_box(0); +x_48 = lean_array_uset(x_18, x_31, x_47); +x_49 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(x_2, x_3, x_32); +x_50 = lean_array_uset(x_48, x_31, x_49); +lean_ctor_set(x_15, 1, x_50); +x_51 = 1; +lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_51); +return x_1; } } else { -lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); -lean_dec(x_1); -x_21 = 5; -x_22 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_23 = lean_usize_land(x_2, x_22); -x_24 = lean_usize_to_nat(x_23); -x_25 = lean_box(2); -x_26 = lean_array_get(x_25, x_20, x_24); -lean_dec(x_24); -lean_dec(x_20); -switch (lean_obj_tag(x_26)) { -case 0: +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; uint64_t x_61; size_t x_62; size_t x_63; size_t x_64; size_t x_65; size_t x_66; lean_object* x_67; uint8_t x_68; +x_52 = lean_ctor_get(x_15, 0); +x_53 = lean_ctor_get(x_15, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_15); +x_54 = lean_array_get_size(x_53); +x_55 = l_Lean_Name_hash___override(x_2); +x_56 = 32; +x_57 = lean_uint64_shift_right(x_55, x_56); +x_58 = lean_uint64_xor(x_55, x_57); +x_59 = 16; +x_60 = lean_uint64_shift_right(x_58, x_59); +x_61 = lean_uint64_xor(x_58, x_60); +x_62 = lean_uint64_to_usize(x_61); +x_63 = lean_usize_of_nat(x_54); +lean_dec(x_54); +x_64 = 1; +x_65 = lean_usize_sub(x_63, x_64); +x_66 = lean_usize_land(x_62, x_65); +x_67 = lean_array_uget(x_53, x_66); +x_68 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_2, x_67); +if (x_68 == 0) { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_name_eq(x_3, x_27); -lean_dec(x_27); -if (x_29 == 0) +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_69 = lean_unsigned_to_nat(1u); +x_70 = lean_nat_add(x_52, x_69); +lean_dec(x_52); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_2); +lean_ctor_set(x_71, 1, x_3); +lean_ctor_set(x_71, 2, x_67); +x_72 = lean_array_uset(x_53, x_66, x_71); +x_73 = lean_unsigned_to_nat(4u); +x_74 = lean_nat_mul(x_70, x_73); +x_75 = lean_unsigned_to_nat(3u); +x_76 = lean_nat_div(x_74, x_75); +lean_dec(x_74); +x_77 = lean_array_get_size(x_72); +x_78 = lean_nat_dec_le(x_76, x_77); +lean_dec(x_77); +lean_dec(x_76); +if (x_78 == 0) { -lean_object* x_30; -lean_dec(x_28); -x_30 = lean_box(0); -return x_30; +lean_object* x_79; lean_object* x_80; uint8_t x_81; +x_79 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_72); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_70); +lean_ctor_set(x_80, 1, x_79); +x_81 = 1; +lean_ctor_set(x_1, 0, x_80); +lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_81); +return x_1; } else { -lean_object* x_31; -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_28); -return x_31; -} +lean_object* x_82; uint8_t x_83; +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_70); +lean_ctor_set(x_82, 1, x_72); +x_83 = 1; +lean_ctor_set(x_1, 0, x_82); +lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_83); +return x_1; } -case 1: -{ -lean_object* x_32; size_t x_33; -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_usize_shift_right(x_2, x_21); -x_1 = x_32; -x_2 = x_33; -goto _start; } -default: +else { -lean_object* x_35; -x_35 = lean_box(0); -return x_35; -} +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_84 = lean_box(0); +x_85 = lean_array_uset(x_53, x_66, x_84); +x_86 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(x_2, x_3, x_67); +x_87 = lean_array_uset(x_85, x_66, x_86); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_52); +lean_ctor_set(x_88, 1, x_87); +x_89 = 1; +lean_ctor_set(x_1, 0, x_88); +lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_89); +return x_1; } } } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_1, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 1); -lean_inc(x_37); +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint64_t x_96; uint64_t x_97; uint64_t x_98; uint64_t x_99; uint64_t x_100; uint64_t x_101; uint64_t x_102; size_t x_103; size_t x_104; size_t x_105; size_t x_106; size_t x_107; lean_object* x_108; uint8_t x_109; +x_90 = lean_ctor_get(x_1, 0); +x_91 = lean_ctor_get(x_1, 1); +lean_inc(x_91); +lean_inc(x_90); lean_dec(x_1); -x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4(x_36, x_37, lean_box(0), x_38, x_3); -lean_dec(x_37); -lean_dec(x_36); -return x_39; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Name_hash___override(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3(x_1, x_4, x_2); -return x_5; -} +x_92 = lean_ctor_get(x_90, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_90, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_94 = x_90; +} else { + lean_dec_ref(x_90); + x_94 = lean_box(0); } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) +x_95 = lean_array_get_size(x_93); +x_96 = l_Lean_Name_hash___override(x_2); +x_97 = 32; +x_98 = lean_uint64_shift_right(x_96, x_97); +x_99 = lean_uint64_xor(x_96, x_98); +x_100 = 16; +x_101 = lean_uint64_shift_right(x_99, x_100); +x_102 = lean_uint64_xor(x_99, x_101); +x_103 = lean_uint64_to_usize(x_102); +x_104 = lean_usize_of_nat(x_95); +lean_dec(x_95); +x_105 = 1; +x_106 = lean_usize_sub(x_104, x_105); +x_107 = lean_usize_land(x_103, x_106); +x_108 = lean_array_uget(x_93, x_107); +x_109 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_2, x_108); +if (x_109 == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_110 = lean_unsigned_to_nat(1u); +x_111 = lean_nat_add(x_92, x_110); +lean_dec(x_92); +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_2); +lean_ctor_set(x_112, 1, x_3); +lean_ctor_set(x_112, 2, x_108); +x_113 = lean_array_uset(x_93, x_107, x_112); +x_114 = lean_unsigned_to_nat(4u); +x_115 = lean_nat_mul(x_111, x_114); +x_116 = lean_unsigned_to_nat(3u); +x_117 = lean_nat_div(x_115, x_116); +lean_dec(x_115); +x_118 = lean_array_get_size(x_113); +x_119 = lean_nat_dec_le(x_117, x_118); +lean_dec(x_118); +lean_dec(x_117); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; +x_120 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_113); +if (lean_is_scalar(x_94)) { + x_121 = lean_alloc_ctor(0, 2, 0); +} else { + x_121 = x_94; +} +lean_ctor_set(x_121, 0, x_111); +lean_ctor_set(x_121, 1, x_120); +x_122 = 1; +x_123 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_91); +lean_ctor_set_uint8(x_123, sizeof(void*)*2, x_122); +return x_123; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_name_eq(x_4, x_1); -if (x_7 == 0) -{ -x_2 = x_6; -goto _start; +lean_object* x_124; uint8_t x_125; lean_object* x_126; +if (lean_is_scalar(x_94)) { + x_124 = lean_alloc_ctor(0, 2, 0); +} else { + x_124 = x_94; +} +lean_ctor_set(x_124, 0, x_111); +lean_ctor_set(x_124, 1, x_113); +x_125 = 1; +x_126 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_91); +lean_ctor_set_uint8(x_126, sizeof(void*)*2, x_125); +return x_126; +} } else { -lean_object* x_9; -lean_inc(x_5); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_5); -return x_9; +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; +x_127 = lean_box(0); +x_128 = lean_array_uset(x_93, x_107, x_127); +x_129 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(x_2, x_3, x_108); +x_130 = lean_array_uset(x_128, x_107, x_129); +if (lean_is_scalar(x_94)) { + x_131 = lean_alloc_ctor(0, 2, 0); +} else { + x_131 = x_94; +} +lean_ctor_set(x_131, 0, x_92); +lean_ctor_set(x_131, 1, x_130); +x_132 = 1; +x_133 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_91); +lean_ctor_set_uint8(x_133, sizeof(void*)*2, x_132); +return x_133; +} } } } } -LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* lean_environment_add(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; -x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +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; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; +lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -lean_dec(x_1); -x_6 = lean_ctor_get(x_4, 1); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_array_get_size(x_6); -x_8 = l_Lean_Name_hash___override(x_2); -x_9 = 32; -x_10 = lean_uint64_shift_right(x_8, x_9); -x_11 = lean_uint64_xor(x_8, x_10); -x_12 = 16; -x_13 = lean_uint64_shift_right(x_11, x_12); -x_14 = lean_uint64_xor(x_11, x_13); -x_15 = lean_uint64_to_usize(x_14); -x_16 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_17 = 1; -x_18 = lean_usize_sub(x_16, x_17); -x_19 = lean_usize_land(x_15, x_18); -x_20 = lean_array_uget(x_6, x_19); -lean_dec(x_6); -x_21 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_2, x_20); -lean_dec(x_20); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; -x_22 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2(x_5, x_2); -return x_22; -} -else -{ -uint8_t x_23; -lean_dec(x_5); -x_23 = !lean_is_exclusive(x_21); -if (x_23 == 0) -{ -return x_21; -} -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 0); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -return x_25; -} -} +x_5 = l_Lean_ConstantInfo_name(x_2); +x_6 = l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__1(x_4, x_5, x_2); +lean_ctor_set(x_1, 0, x_6); +return x_1; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; -x_26 = lean_ctor_get(x_1, 0); -lean_inc(x_26); +lean_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; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_ctor_get(x_1, 2); +x_11 = lean_ctor_get(x_1, 3); +x_12 = lean_ctor_get(x_1, 4); +x_13 = lean_ctor_get(x_1, 5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_7); lean_dec(x_1); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_array_get_size(x_27); -x_29 = l_Lean_Name_hash___override(x_2); -x_30 = 32; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = 16; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = lean_uint64_to_usize(x_35); -x_37 = lean_usize_of_nat(x_28); -lean_dec(x_28); -x_38 = 1; -x_39 = lean_usize_sub(x_37, x_38); -x_40 = lean_usize_land(x_36, x_39); -x_41 = lean_array_uget(x_27, x_40); -lean_dec(x_27); -x_42 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_2, x_41); -lean_dec(x_41); -return x_42; +x_14 = l_Lean_ConstantInfo_name(x_2); +x_15 = l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__1(x_7, x_14, x_2); +x_16 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_9); +lean_ctor_set(x_16, 2, x_10); +lean_ctor_set(x_16, 3, x_11); +lean_ctor_set(x_16, 4, x_12); +lean_ctor_set(x_16, 5, x_13); +lean_ctor_set_uint8(x_16, sizeof(void*)*6, x_8); +return x_16; } } } -LEAN_EXPORT lean_object* lean_environment_find(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_3; lean_object* x_4; -x_3 = lean_ctor_get(x_1, 1); -lean_inc(x_3); +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); lean_dec(x_1); -x_4 = l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(x_3, x_2); -lean_dec(x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__4(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_6; +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Environment_find_x3f___spec__3(x_1, x_4, x_3); +x_7 = lean_unbox_usize(x_3); lean_dec(x_3); -return x_5; +x_8 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3(x_1, x_6, x_7, x_4, x_5); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2(x_1, x_2); +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_1, x_2); lean_dec(x_2); -return x_3; +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t lean_kernel_diag_is_enabled(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_1, x_2); -lean_dec(x_2); +uint8_t x_2; +x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); lean_dec(x_1); -return x_3; +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_Diagnostics_isEnabled___boxed(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(x_1, x_2); -lean_dec(x_2); +uint8_t x_2; lean_object* x_3; +x_2 = lean_kernel_diag_is_enabled(x_1); +x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT uint8_t l_Lean_Environment_contains(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object* x_1, uint8_t x_2) { _start: { -lean_object* x_3; uint8_t x_4; -x_3 = lean_ctor_get(x_1, 1); -lean_inc(x_3); -lean_dec(x_1); -x_4 = l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(x_3, x_2); -return x_4; -} +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_ctor_set_uint8(x_4, sizeof(void*)*1, x_2); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_Environment_contains___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Environment_contains(x_1, x_2); -lean_dec(x_2); -x_4 = lean_box(x_3); -return x_4; +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set_uint8(x_7, sizeof(void*)*1, x_2); +lean_ctor_set(x_1, 1, x_7); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Environment_imports(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; -x_2 = lean_ctor_get(x_1, 4); -x_3 = lean_ctor_get(x_2, 1); -lean_inc(x_3); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Environment_imports___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Environment_imports(x_1); +lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_1, 3); +x_13 = lean_ctor_get(x_1, 4); +x_14 = lean_ctor_get(x_1, 5); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_8); +lean_inc(x_9); lean_dec(x_1); -return x_2; -} +x_15 = lean_ctor_get(x_8, 0); +lean_inc(x_15); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + x_16 = x_8; +} else { + lean_dec_ref(x_8); + x_16 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_Environment_allImportedModuleNames(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = lean_ctor_get(x_1, 4); -x_3 = lean_ctor_get(x_2, 3); -lean_inc(x_3); -return x_3; +if (lean_is_scalar(x_16)) { + x_17 = lean_alloc_ctor(0, 1, 1); +} else { + x_17 = x_16; } +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_2); +x_18 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set(x_18, 2, x_11); +lean_ctor_set(x_18, 3, x_12); +lean_ctor_set(x_18, 4, x_13); +lean_ctor_set(x_18, 5, x_14); +lean_ctor_set_uint8(x_18, sizeof(void*)*6, x_10); +return x_18; } -LEAN_EXPORT lean_object* l_Lean_Environment_allImportedModuleNames___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Environment_allImportedModuleNames(x_1); -lean_dec(x_1); -return x_2; } } -LEAN_EXPORT lean_object* lean_environment_set_main_module(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_enableDiag___boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_ctor_get(x_1, 4); -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) -{ -lean_object* x_6; -x_6 = lean_ctor_get(x_4, 0); -lean_dec(x_6); -lean_ctor_set(x_4, 0, x_2); -return x_1; -} -else -{ -uint32_t x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_7 = lean_ctor_get_uint32(x_4, sizeof(void*)*5); -x_8 = lean_ctor_get_uint8(x_4, sizeof(void*)*5 + 4); -x_9 = lean_ctor_get(x_4, 1); -x_10 = lean_ctor_get(x_4, 2); -x_11 = lean_ctor_get(x_4, 3); -x_12 = lean_ctor_get(x_4, 4); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_4); -x_13 = lean_alloc_ctor(0, 5, 5); -lean_ctor_set(x_13, 0, x_2); -lean_ctor_set(x_13, 1, x_9); -lean_ctor_set(x_13, 2, x_10); -lean_ctor_set(x_13, 3, x_11); -lean_ctor_set(x_13, 4, x_12); -lean_ctor_set_uint32(x_13, sizeof(void*)*5, x_7); -lean_ctor_set_uint8(x_13, sizeof(void*)*5 + 4, x_8); -lean_ctor_set(x_1, 4, x_13); -return x_1; +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l_Lean_Kernel_Environment_enableDiag(x_1, x_3); +return x_4; } } -else +LEAN_EXPORT uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object* x_1) { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint32_t x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_14 = lean_ctor_get(x_1, 4); -x_15 = lean_ctor_get(x_1, 0); -x_16 = lean_ctor_get(x_1, 1); -x_17 = lean_ctor_get(x_1, 2); -x_18 = lean_ctor_get(x_1, 3); -lean_inc(x_14); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_1); -x_19 = lean_ctor_get_uint32(x_14, sizeof(void*)*5); -x_20 = lean_ctor_get_uint8(x_14, sizeof(void*)*5 + 4); -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); -x_22 = lean_ctor_get(x_14, 2); -lean_inc(x_22); -x_23 = lean_ctor_get(x_14, 3); -lean_inc(x_23); -x_24 = lean_ctor_get(x_14, 4); -lean_inc(x_24); -if (lean_is_exclusive(x_14)) { - lean_ctor_release(x_14, 0); - lean_ctor_release(x_14, 1); - lean_ctor_release(x_14, 2); - lean_ctor_release(x_14, 3); - lean_ctor_release(x_14, 4); - x_25 = x_14; -} else { - lean_dec_ref(x_14); - x_25 = lean_box(0); -} -if (lean_is_scalar(x_25)) { - x_26 = lean_alloc_ctor(0, 5, 5); -} else { - x_26 = x_25; -} -lean_ctor_set(x_26, 0, x_2); -lean_ctor_set(x_26, 1, x_21); -lean_ctor_set(x_26, 2, x_22); -lean_ctor_set(x_26, 3, x_23); -lean_ctor_set(x_26, 4, x_24); -lean_ctor_set_uint32(x_26, sizeof(void*)*5, x_19); -lean_ctor_set_uint8(x_26, sizeof(void*)*5 + 4, x_20); -x_27 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_16); -lean_ctor_set(x_27, 2, x_17); -lean_ctor_set(x_27, 3, x_18); -lean_ctor_set(x_27, 4, x_26); -return x_27; -} +lean_object* x_2; uint8_t x_3; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get_uint8(x_2, sizeof(void*)*1); +return x_3; } } -LEAN_EXPORT lean_object* lean_environment_main_module(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_isDiagnosticsEnabled___boxed(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = lean_ctor_get(x_1, 4); -lean_inc(x_2); +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_1); lean_dec(x_1); -x_3 = lean_ctor_get(x_2, 0); -lean_inc(x_3); -lean_dec(x_2); +x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* lean_environment_mark_quot_init(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Kernel_Environment_resetDiag(lean_object* x_1) { _start: { uint8_t x_2; @@ -3430,5166 +3220,4664 @@ x_2 = !lean_is_exclusive(x_1); if (x_2 == 0) { lean_object* x_3; uint8_t x_4; -x_3 = lean_ctor_get(x_1, 4); +x_3 = lean_ctor_get(x_1, 1); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) { -uint8_t x_5; -x_5 = 1; -lean_ctor_set_uint8(x_3, sizeof(void*)*5 + 4, x_5); +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 0); +lean_dec(x_5); +x_6 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +lean_ctor_set(x_3, 0, x_6); return x_1; } else { -uint32_t 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_6 = lean_ctor_get_uint32(x_3, sizeof(void*)*5); -x_7 = lean_ctor_get(x_3, 0); -x_8 = lean_ctor_get(x_3, 1); -x_9 = lean_ctor_get(x_3, 2); -x_10 = lean_ctor_get(x_3, 3); -x_11 = lean_ctor_get(x_3, 4); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); +uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); lean_dec(x_3); -x_12 = 1; -x_13 = lean_alloc_ctor(0, 5, 5); -lean_ctor_set(x_13, 0, x_7); -lean_ctor_set(x_13, 1, x_8); -lean_ctor_set(x_13, 2, x_9); -lean_ctor_set(x_13, 3, x_10); -lean_ctor_set(x_13, 4, x_11); -lean_ctor_set_uint32(x_13, sizeof(void*)*5, x_6); -lean_ctor_set_uint8(x_13, sizeof(void*)*5 + 4, x_12); -lean_ctor_set(x_1, 4, x_13); +x_8 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +x_9 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_7); +lean_ctor_set(x_1, 1, x_9); return x_1; } } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint32_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; -x_14 = lean_ctor_get(x_1, 4); -x_15 = lean_ctor_get(x_1, 0); -x_16 = lean_ctor_get(x_1, 1); -x_17 = lean_ctor_get(x_1, 2); -x_18 = lean_ctor_get(x_1, 3); -lean_inc(x_14); -lean_inc(x_18); -lean_inc(x_17); +lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_13 = lean_ctor_get(x_1, 2); +x_14 = lean_ctor_get(x_1, 3); +x_15 = lean_ctor_get(x_1, 4); +x_16 = lean_ctor_get(x_1, 5); lean_inc(x_16); lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_10); +lean_inc(x_11); lean_dec(x_1); -x_19 = lean_ctor_get_uint32(x_14, sizeof(void*)*5); -x_20 = lean_ctor_get(x_14, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); -x_22 = lean_ctor_get(x_14, 2); -lean_inc(x_22); -x_23 = lean_ctor_get(x_14, 3); -lean_inc(x_23); -x_24 = lean_ctor_get(x_14, 4); -lean_inc(x_24); -if (lean_is_exclusive(x_14)) { - lean_ctor_release(x_14, 0); - lean_ctor_release(x_14, 1); - lean_ctor_release(x_14, 2); - lean_ctor_release(x_14, 3); - lean_ctor_release(x_14, 4); - x_25 = x_14; +x_17 = lean_ctor_get_uint8(x_10, sizeof(void*)*1); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_18 = x_10; } else { - lean_dec_ref(x_14); - x_25 = lean_box(0); + lean_dec_ref(x_10); + x_18 = lean_box(0); } -x_26 = 1; -if (lean_is_scalar(x_25)) { - x_27 = lean_alloc_ctor(0, 5, 5); +x_19 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +if (lean_is_scalar(x_18)) { + x_20 = lean_alloc_ctor(0, 1, 1); } else { - x_27 = x_25; -} -lean_ctor_set(x_27, 0, x_20); -lean_ctor_set(x_27, 1, x_21); -lean_ctor_set(x_27, 2, x_22); -lean_ctor_set(x_27, 3, x_23); -lean_ctor_set(x_27, 4, x_24); -lean_ctor_set_uint32(x_27, sizeof(void*)*5, x_19); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 4, x_26); -x_28 = lean_alloc_ctor(0, 5, 0); -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_27); -return x_28; -} + x_20 = x_18; } +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set_uint8(x_20, sizeof(void*)*1, x_17); +x_21 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_21, 0, x_11); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_13); +lean_ctor_set(x_21, 3, x_14); +lean_ctor_set(x_21, 4, x_15); +lean_ctor_set(x_21, 5, x_16); +lean_ctor_set_uint8(x_21, sizeof(void*)*6, x_12); +return x_21; } -LEAN_EXPORT uint8_t lean_environment_quot_init(lean_object* x_1) { -_start: -{ -lean_object* x_2; uint8_t x_3; -x_2 = lean_ctor_get(x_1, 4); -lean_inc(x_2); -lean_dec(x_1); -x_3 = lean_ctor_get_uint8(x_2, sizeof(void*)*5 + 4); -lean_dec(x_2); -return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_isQuotInit___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_2; lean_object* x_3; -x_2 = lean_environment_quot_init(x_1); -x_3 = lean_box(x_2); -return x_3; +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; } +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_name_eq(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; } -LEAN_EXPORT uint32_t lean_environment_trust_level(lean_object* x_1) { -_start: +else { -lean_object* x_2; uint32_t x_3; -x_2 = lean_ctor_get(x_1, 4); -lean_inc(x_2); -lean_dec(x_1); -x_3 = lean_ctor_get_uint32(x_2, sizeof(void*)*5); -lean_dec(x_2); -return x_3; +lean_object* x_14; lean_object* x_15; +x_14 = lean_array_fget(x_2, x_4); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_getTrustLevel___boxed(lean_object* x_1) { -_start: -{ -uint32_t x_2; lean_object* x_3; -x_2 = lean_environment_trust_level(x_1); -x_3 = lean_box_uint32(x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_2) == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; -} -else +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_name_eq(x_4, x_1); -if (x_7 == 0) +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); +lean_dec(x_5); +switch (lean_obj_tag(x_11)) { +case 0: { -x_2 = x_6; -goto _start; +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_name_eq(x_3, x_12); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; } else { -lean_object* x_9; -lean_inc(x_5); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_5); -return x_9; -} +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); +return x_1; } } +case 1: +{ +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object* x_1, lean_object* x_2) { -_start: +default: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; -x_3 = lean_ctor_get(x_1, 0); -x_4 = lean_ctor_get(x_3, 1); -x_5 = lean_array_get_size(x_4); -x_6 = l_Lean_Name_hash___override(x_2); -x_7 = 32; -x_8 = lean_uint64_shift_right(x_6, x_7); -x_9 = lean_uint64_xor(x_6, x_8); -x_10 = 16; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = lean_uint64_to_usize(x_12); -x_14 = lean_usize_of_nat(x_5); -lean_dec(x_5); -x_15 = 1; -x_16 = lean_usize_sub(x_14, x_15); -x_17 = lean_usize_land(x_13, x_16); -x_18 = lean_array_uget(x_4, x_17); -x_19 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1(x_2, x_18); -lean_dec(x_18); +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); return x_19; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} } -LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_Lean_Environment_getModuleIdxFor_x3f(x_1, x_2); -lean_dec(x_2); +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT uint8_t l_Lean_Environment_isConstructor(lean_object* x_1, lean_object* x_2) { -_start: +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { +case 0: { -lean_object* x_3; -x_3 = lean_environment_find(x_1, x_2); -if (lean_obj_tag(x_3) == 0) +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_name_eq(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) { -uint8_t x_4; -x_4 = 0; -return x_4; +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; } else { -lean_object* x_5; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -lean_dec(x_3); -if (lean_obj_tag(x_5) == 6) +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; +} +} +case 1: { -uint8_t x_6; -lean_dec(x_5); -x_6 = 1; -return x_6; +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; } -else +default: { -uint8_t x_7; -lean_dec(x_5); -x_7 = 0; -return x_7; +lean_object* x_35; +x_35 = lean_box(0); +return x_35; } } } } -LEAN_EXPORT lean_object* l_Lean_Environment_isConstructor___boxed(lean_object* x_1, lean_object* x_2) { +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__3(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Environment_isConstructor(x_1, x_2); -x_4 = lean_box(x_3); -return x_4; +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__2(x_1, x_4, x_2); +return x_5; } } -LEAN_EXPORT uint8_t l_Lean_Environment_isSafeDefinition(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__6(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_3; -x_3 = lean_environment_find(x_1, x_2); -if (lean_obj_tag(x_3) == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) { -uint8_t x_4; -x_4 = 0; -return x_4; +lean_dec(x_5); +return x_6; } else { -lean_object* x_5; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -lean_dec(x_3); -if (lean_obj_tag(x_5) == 1) +lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Name_hash___override(x_9); +x_12 = lean_uint64_to_usize(x_11); +x_13 = 1; +x_14 = lean_usize_sub(x_1, x_13); +x_15 = 5; +x_16 = lean_usize_mul(x_15, x_14); +x_17 = lean_usize_shift_right(x_12, x_16); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_5, x_18); +lean_dec(x_5); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5(x_6, x_17, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_19; +x_6 = x_20; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_6; uint8_t x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_5, 0); +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_ctor_get_uint8(x_6, sizeof(void*)*4); -lean_dec(x_6); -x_8 = lean_box(x_7); -if (lean_obj_tag(x_8) == 1) +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) { uint8_t x_9; -x_9 = 1; -return x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; } else { -uint8_t x_10; -lean_dec(x_8); -x_10 = 0; -return x_10; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } else { -uint8_t x_11; +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_name_eq(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_6); lean_dec(x_5); -x_11 = 0; -return x_11; -} -} -} +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_2); +x_2 = x_20; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Environment_isSafeDefinition___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Environment_isSafeDefinition(x_1, x_2); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT uint8_t l_Lean_Environment_getModuleIdx_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = lean_name_eq(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_3 = lean_alloc_closure((void*)(l_Lean_Environment_getModuleIdx_x3f___lambda__1___boxed), 2, 1); -lean_closure_set(x_3, 0, x_2); -x_4 = lean_ctor_get(x_1, 4); -x_5 = lean_ctor_get(x_4, 3); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_findIdx_x3f_loop___rarg(x_3, x_5, x_6); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 0) { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Environment_getModuleIdx_x3f___lambda__1(x_1, x_2); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_1, 0); +lean_dec(x_24); +x_25 = lean_array_fset(x_5, x_2, x_3); +x_26 = lean_array_fset(x_6, x_2, x_4); lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_Lean_Environment_getModuleIdx_x3f(x_1, x_2); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Environment_addDeclCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; lean_object* x_6; -x_5 = lean_unbox_usize(x_2); +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); lean_dec(x_2); -x_6 = lean_add_decl(x_1, x_5, x_3, x_4); -lean_dec(x_4); -lean_dec(x_3); -return x_6; -} +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } -LEAN_EXPORT lean_object* l_Lean_Environment_addDeclWithoutChecking___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_add_decl_without_checking(x_1, x_2); -lean_dec(x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_ConstantInfo_type(x_1); -x_4 = l_Lean_ConstantInfo_levelParams(x_1); -x_5 = l_Lean_Expr_instantiateLevelParams(x_3, x_4, x_2); -lean_dec(x_3); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_3; -x_3 = l_Lean_ConstantInfo_instantiateTypeLevelParams(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21(lean_object* x_1, lean_object* x_2) { -_start: +if (lean_obj_tag(x_1) == 0) { -uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = 0; -x_4 = l_Lean_ConstantInfo_value_x21(x_1, x_3); -x_5 = l_Lean_ConstantInfo_levelParams(x_1); -x_6 = l_Lean_Expr_instantiateLevelParams(x_4, x_5, x_2); -lean_dec(x_4); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21___boxed(lean_object* x_1, lean_object* x_2) { -_start: +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) { -lean_object* x_3; -x_3 = l_Lean_ConstantInfo_instantiateValueLevelParams_x21(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) { -lean_inc(x_2); -return x_2; -} +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = lean_apply_1(x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -lean_inc(x_3); -return x_3; -} -} -static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__1() { -_start: +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__2() { -_start: +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_name_eq(x_4, x_19); +if (x_21 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_2 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_array_fset(x_17, x_12, x_23); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_24); +return x_1; } -static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__1___boxed), 2, 0); +lean_object* x_25; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_25 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_25); return x_1; } } -static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__4() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__2), 3, 0); +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_name_eq(x_4, x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = lean_array_fset(x_17, x_12, x_30); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_31); return x_1; } -} -static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__3___boxed), 4, 0); +lean_object* x_32; lean_object* x_33; +lean_dec(x_27); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_5); +x_33 = lean_array_fset(x_17, x_12, x_32); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_33); return x_1; } } -static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__6() { -_start: +} +case 1: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 0); +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) +{ +lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_usize_shift_right(x_2, x_9); +x_37 = lean_usize_add(x_3, x_8); +x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5(x_35, x_36, x_37, x_4, x_5); +lean_ctor_set(x_15, 0, x_38); +x_39 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); return x_1; } -} -static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_instInhabitedEnvExtensionInterface___closed__3; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__4; -x_3 = l_Lean_instInhabitedEnvExtensionInterface___closed__5; -x_4 = l_Lean_instInhabitedEnvExtensionInterface___closed__2; -x_5 = l_Lean_instInhabitedEnvExtensionInterface___closed__6; -x_6 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_6, 0, x_1); -lean_ctor_set(x_6, 1, x_2); -lean_ctor_set(x_6, 2, x_3); -lean_ctor_set(x_6, 3, x_3); -lean_ctor_set(x_6, 4, x_3); -lean_ctor_set(x_6, 5, x_4); -lean_ctor_set(x_6, 6, x_5); -return x_6; +lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_15, 0); +lean_inc(x_40); +lean_dec(x_15); +x_41 = lean_usize_shift_right(x_2, x_9); +x_42 = lean_usize_add(x_3, x_8); +x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5(x_40, x_41, x_42, x_4, x_5); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_array_fset(x_17, x_12, x_44); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_45); +return x_1; } } -static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface() { -_start: +default: { -lean_object* x_1; -x_1 = l_Lean_instInhabitedEnvExtensionInterface___closed__7; +lean_object* x_46; lean_object* x_47; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_4); +lean_ctor_set(x_46, 1, x_5); +x_47 = lean_array_fset(x_17, x_12, x_46); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_47); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_instInhabitedEnvExtensionInterface___lambda__1(x_1, x_2); -lean_dec(x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_instInhabitedEnvExtensionInterface___lambda__3(x_1, x_2, x_3, x_4); +lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +lean_dec(x_1); +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_52 = lean_usize_land(x_2, x_51); +x_53 = lean_usize_to_nat(x_52); +x_54 = lean_array_get_size(x_48); +x_55 = lean_nat_dec_lt(x_53, x_54); +lean_dec(x_54); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_53); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; -} +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; } -static uint32_t _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; uint32_t x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_uint32_of_nat(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__2() { -_start: +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_array_fget(x_48, x_53); +x_58 = lean_box(0); +x_59 = lean_array_fset(x_48, x_53, x_58); +switch (lean_obj_tag(x_57)) { +case 0: { -lean_object* x_1; uint32_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(0); -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1; -x_3 = l_Lean_instToStringImport___closed__2; -x_4 = lean_alloc_ctor(0, 2, 4); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_3); -lean_ctor_set_uint32(x_4, sizeof(void*)*2, x_2); -return x_4; +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_62 = x_57; +} else { + lean_dec_ref(x_57); + x_62 = lean_box(0); } +x_63 = lean_name_eq(x_4, x_60); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_66 = lean_array_fset(x_59, x_53, x_65); +lean_dec(x_53); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_61); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_62; +} +lean_ctor_set(x_68, 0, x_4); +lean_ctor_set(x_68, 1, x_5); +x_69 = lean_array_fset(x_59, x_53, x_68); +lean_dec(x_53); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1() { -_start: +case 1: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1), 1, 0); -return x_1; +lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_71 = lean_ctor_get(x_57, 0); +lean_inc(x_71); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_72 = x_57; +} else { + lean_dec_ref(x_57); + x_72 = lean_box(0); } +x_73 = lean_usize_shift_right(x_2, x_50); +x_74 = lean_usize_add(x_3, x_49); +x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5(x_71, x_73, x_74, x_4, x_5); +if (lean_is_scalar(x_72)) { + x_76 = lean_alloc_ctor(1, 1, 0); +} else { + x_76 = x_72; } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2() { -_start: +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_array_fset(x_59, x_53, x_76); +lean_dec(x_53); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; +} +default: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_4); +lean_ctor_set(x_79, 1, x_5); +x_80 = lean_array_fset(x_59, x_53, x_79); +lean_dec(x_53); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1472_(lean_object* x_1) { -_start: +} +else { -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_3 = lean_st_mk_ref(x_2, x_1); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) { -return x_3; +lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; +x_83 = lean_unsigned_to_nat(0u); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__7(x_1, x_83, x_4, x_5); +x_85 = 7; +x_86 = lean_usize_dec_le(x_85, x_3); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_dec_lt(x_87, x_88); +lean_dec(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_84, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 1); +lean_inc(x_91); +lean_dec(x_84); +x_92 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__6(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_6); -lean_inc(x_5); -lean_dec(x_3); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_5); -lean_ctor_set(x_7, 1, x_6); -return x_7; -} +return x_84; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef; -return x_1; +return x_84; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; -x_5 = lean_st_ref_get(x_4, x_3); -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_7 = lean_ctor_get(x_5, 0); -x_8 = lean_ctor_get(x_5, 1); -x_9 = lean_array_get_size(x_7); -x_10 = lean_nat_dec_lt(x_1, x_9); -lean_dec(x_9); -if (x_10 == 0) -{ -lean_dec(x_7); +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); +lean_inc(x_94); lean_dec(x_1); -lean_ctor_set(x_5, 0, x_2); -return x_5; -} -else +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_unsigned_to_nat(0u); +x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__7(x_96, x_97, x_4, x_5); +x_99 = 7; +x_100 = lean_usize_dec_le(x_99, x_3); +if (x_100 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_free_object(x_5); -x_11 = lean_array_fget(x_7, x_1); -lean_dec(x_7); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_apply_1(x_12, x_8); -if (lean_obj_tag(x_13) == 0) +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_dec_lt(x_101, x_102); +lean_dec(x_101); +if (x_103 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_array_push(x_2, x_14); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_add(x_1, x_17); -lean_dec(x_1); -x_1 = x_18; -x_2 = x_16; -x_3 = x_15; -goto _start; +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__6(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; } else { -uint8_t x_20; -lean_dec(x_2); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_13); -if (x_20 == 0) -{ -return x_13; +return x_98; +} } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_13, 0); -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_13); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +return x_98; } } } } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_ctor_get(x_5, 0); -x_25 = lean_ctor_get(x_5, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_5); -x_26 = lean_array_get_size(x_24); -x_27 = lean_nat_dec_lt(x_1, x_26); -lean_dec(x_26); -if (x_27 == 0) +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Name_hash___override(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5(x_1, x_5, x_6, x_2, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* lean_kernel_record_unfold(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_28; -lean_dec(x_24); -lean_dec(x_1); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_2); -lean_ctor_set(x_28, 1, x_25); -return x_28; +uint8_t x_3; +x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); +if (x_3 == 0) +{ +lean_dec(x_2); +return x_1; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_array_fget(x_24, x_1); -lean_dec(x_24); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = lean_apply_1(x_30, x_25); -if (lean_obj_tag(x_31) == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_array_push(x_2, x_32); -x_35 = lean_unsigned_to_nat(1u); -x_36 = lean_nat_add(x_1, x_35); -lean_dec(x_1); -x_1 = x_36; -x_2 = x_34; -x_3 = x_33; -goto _start; +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_5, x_2); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_unsigned_to_nat(1u); +x_8 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_5, x_2, x_7); +lean_ctor_set(x_1, 0, x_8); +return x_1; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -lean_dec(x_2); -lean_dec(x_1); -x_38 = lean_ctor_get(x_31, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_31, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_40 = x_31; -} else { - lean_dec_ref(x_31); - x_40 = lean_box(0); -} -if (lean_is_scalar(x_40)) { - x_41 = lean_alloc_ctor(1, 2, 0); -} else { - x_41 = x_40; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_9, x_10); +lean_dec(x_9); +x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_5, x_2, x_11); +lean_ctor_set(x_1, 0, x_12); +return x_1; } -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set(x_41, 1, x_39); -return x_41; } +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +lean_inc(x_13); +x_14 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_13, x_2); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_unsigned_to_nat(1u); +x_16 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_13, x_2, x_15); +x_17 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_3); +return x_17; } +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_14, 0); +lean_inc(x_18); +lean_dec(x_14); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_18, x_19); +lean_dec(x_18); +x_21 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_13, x_2, x_20); +x_22 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set_uint8(x_22, sizeof(void*)*1, x_3); +return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; -x_3 = lean_array_get_size(x_1); -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop(x_3, x_1, x_2); -return x_4; } } -static lean_object* _init_l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid environment extension has been accessed", 47, 47); -return x_1; +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_6; } } -static lean_object* _init_l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg___closed__1; -return x_1; +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_setState___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_panic_fn(x_1, x_2); +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_1, x_2); +lean_dec(x_2); return x_3; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Environment", 16, 16); -return x_1; +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__6(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__2() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.EnvExtensionInterfaceUnsafe.setState", 41, 41); -return x_1; -} -} -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___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_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__2; -x_3 = lean_unsigned_to_nat(334u); -x_4 = lean_unsigned_to_nat(4u); -x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; -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_EnvExtensionInterfaceUnsafe_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_array_get_size(x_2); -x_6 = lean_nat_dec_lt(x_4, x_5); -lean_dec(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__3; -x_8 = l_panic___at_Lean_EnvExtensionInterfaceUnsafe_setState___spec__1(x_2, x_7); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__5(x_1, x_6, x_7, x_4, x_5); return x_8; } -else -{ -lean_object* x_9; -x_9 = lean_array_fset(x_2, x_4, x_3); -return x_9; -} } -} -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState(lean_object* x_1) { +LEAN_EXPORT lean_object* lean_kernel_get_diag(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___boxed), 3, 0); +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* lean_kernel_set_diag(lean_object* x_1, lean_object* x_2) { _start: { +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_2, x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_dec(x_4); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +else +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_7 = lean_ctor_get(x_1, 2); +x_8 = lean_ctor_get(x_1, 3); +x_9 = lean_ctor_get(x_1, 4); +x_10 = lean_ctor_get(x_1, 5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_5); lean_dec(x_1); -return x_4; +x_11 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_11, 0, x_5); +lean_ctor_set(x_11, 1, x_2); +lean_ctor_set(x_11, 2, x_7); +lean_ctor_set(x_11, 3, x_8); +lean_ctor_set(x_11, 4, x_9); +lean_ctor_set(x_11, 5, x_10); +lean_ctor_set_uint8(x_11, sizeof(void*)*6, x_6); +return x_11; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_modifyState___spec__1(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* lean_elab_environment_of_kernel_env(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = lean_panic_fn(x_1, x_2); -return x_3; +return x_1; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__1() { +LEAN_EXPORT lean_object* lean_elab_environment_to_kernel_env(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.EnvExtensionInterfaceUnsafe.modifyState", 44, 44); return x_1; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Environment_addDeclCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__1; -x_3 = lean_unsigned_to_nat(344u); -x_4 = lean_unsigned_to_nat(4u); -x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +size_t x_5; lean_object* x_6; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_elab_add_decl(x_1, x_5, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Environment_addDeclWithoutChecking___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_array_get_size(x_2); -x_6 = lean_nat_dec_lt(x_4, x_5); -lean_dec(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -lean_dec(x_3); -x_7 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2; -x_8 = l_panic___at_Lean_EnvExtensionInterfaceUnsafe_modifyState___spec__1(x_2, x_7); -return x_8; +lean_object* x_3; +x_3 = lean_elab_add_decl_without_checking(x_1, x_2); +lean_dec(x_2); +return x_3; } -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_array_fget(x_2, x_4); -x_10 = lean_box(0); -x_11 = lean_array_fset(x_2, x_4, x_10); -x_12 = lean_apply_1(x_3, x_9); -x_13 = lean_array_fset(x_11, x_4, x_12); -return x_13; } +LEAN_EXPORT lean_object* l_Lean_Environment_constants(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Environment_constants___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___boxed), 3, 0); +x_2 = l_Lean_Environment_constants(x_1); +lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Environment_const2ModIdx(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 2); +lean_inc(x_2); +return x_2; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Environment_const2ModIdx___boxed(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.EnvExtensionInterfaceUnsafe.getState", 41, 41); -return x_1; +lean_object* x_2; +x_2 = l_Lean_Environment_const2ModIdx(x_1); +lean_dec(x_1); +return x_2; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__2() { +LEAN_EXPORT lean_object* lean_elab_environment_add(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1; -x_3 = lean_unsigned_to_nat(351u); -x_4 = lean_unsigned_to_nat(4u); -x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_3; +x_3 = lean_environment_add(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_array_get_size(x_3); -x_6 = lean_nat_dec_lt(x_4, x_5); -lean_dec(x_5); -if (x_6 == 0) +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) { -lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__2; -x_8 = l_panic___rarg(x_1, x_7); +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; return x_8; } else { -lean_object* x_9; -lean_dec(x_1); -x_9 = lean_array_fget(x_3, x_4); -return x_9; -} -} +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_name_eq(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___boxed), 3, 0); -return x_2; +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -return x_4; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; -x_5 = lean_st_ref_get(x_4, x_3); -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_7 = lean_ctor_get(x_5, 0); -x_8 = lean_ctor_get(x_5, 1); -x_9 = lean_array_get_size(x_7); -lean_dec(x_7); -lean_ctor_set(x_5, 1, x_1); -lean_ctor_set(x_5, 0, x_9); -x_10 = lean_st_ref_take(x_4, x_8); +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; uint8_t x_12; x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); lean_dec(x_10); -lean_inc(x_5); -x_13 = lean_array_push(x_11, x_5); -x_14 = lean_st_ref_set(x_4, x_13, x_12); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; -x_16 = lean_ctor_get(x_14, 0); -lean_dec(x_16); -lean_ctor_set(x_14, 0, x_5); -return x_14; +x_12 = lean_name_eq(x_3, x_11); +lean_dec(x_11); +return x_12; } -else +case 1: { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_5); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_13; size_t x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_usize_shift_right(x_2, x_5); +x_1 = x_13; +x_2 = x_14; +goto _start; +} +default: +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} } } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_19 = lean_ctor_get(x_5, 0); -x_20 = lean_ctor_get(x_5, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_5); -x_21 = lean_array_get_size(x_19); -lean_dec(x_19); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_1); -x_23 = lean_st_ref_take(x_4, x_20); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -lean_inc(x_22); -x_26 = lean_array_push(x_24, x_22); -x_27 = lean_st_ref_set(x_4, x_26, x_25); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - x_29 = x_27; -} else { - lean_dec_ref(x_27); - x_29 = lean_box(0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 1); +lean_inc(x_18); +lean_dec(x_1); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4(x_17, x_18, lean_box(0), x_19, x_3); +lean_dec(x_18); +lean_dec(x_17); +return x_20; } -if (lean_is_scalar(x_29)) { - x_30 = lean_alloc_ctor(0, 2, 0); -} else { - x_30 = x_29; } -lean_ctor_set(x_30, 0, x_22); -lean_ctor_set(x_30, 1, x_28); -return x_30; } +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3(x_1, x_4, x_2); +return x_5; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__1() { +LEAN_EXPORT uint8_t l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("failed to register environment, extensions can only be registered during initialization", 87, 87); -return x_1; +uint8_t x_3; +x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; lean_object* x_20; uint8_t x_21; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_array_get_size(x_6); +x_8 = l_Lean_Name_hash___override(x_2); +x_9 = 32; +x_10 = lean_uint64_shift_right(x_8, x_9); +x_11 = lean_uint64_xor(x_8, x_10); +x_12 = 16; +x_13 = lean_uint64_shift_right(x_11, x_12); +x_14 = lean_uint64_xor(x_11, x_13); +x_15 = lean_uint64_to_usize(x_14); +x_16 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_17 = 1; +x_18 = lean_usize_sub(x_16, x_17); +x_19 = lean_usize_land(x_15, x_18); +x_20 = lean_array_uget(x_6, x_19); +lean_dec(x_6); +x_21 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_2, x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; +x_22 = l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(x_5, x_2); +return x_22; } +else +{ +uint8_t x_23; +lean_dec(x_5); +x_23 = 1; +return x_23; } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__1; -x_2 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; size_t x_38; lean_object* x_39; uint8_t x_40; +x_24 = lean_ctor_get(x_1, 0); +lean_inc(x_24); +lean_dec(x_1); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_array_get_size(x_25); +x_27 = l_Lean_Name_hash___override(x_2); +x_28 = 32; +x_29 = lean_uint64_shift_right(x_27, x_28); +x_30 = lean_uint64_xor(x_27, x_29); +x_31 = 16; +x_32 = lean_uint64_shift_right(x_30, x_31); +x_33 = lean_uint64_xor(x_30, x_32); +x_34 = lean_uint64_to_usize(x_33); +x_35 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_36 = 1; +x_37 = lean_usize_sub(x_35, x_36); +x_38 = lean_usize_land(x_34, x_37); +x_39 = lean_array_uget(x_25, x_38); +lean_dec(x_25); +x_40 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_2, x_39); +lean_dec(x_39); +return x_40; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* l_Lean_Environment_addExtraName(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = l_Lean_initializing(x_2); -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_unbox(x_4); -lean_dec(x_4); -if (x_5 == 0) +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(x_3, x_2); +if (x_4 == 0) { -uint8_t x_6; -lean_dec(x_1); -x_6 = !lean_is_exclusive(x_3); -if (x_6 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_1); +if (x_5 == 0) { -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_3, 0); -lean_dec(x_7); -x_8 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2; -lean_ctor_set_tag(x_3, 1); -lean_ctor_set(x_3, 0, x_8); -return x_3; +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_1, 4); +x_7 = lean_box(0); +x_8 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_6, x_2, x_7); +lean_ctor_set(x_1, 4, x_8); +return x_1; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_3, 1); +lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_ctor_get(x_1, 2); +x_13 = lean_ctor_get(x_1, 3); +x_14 = lean_ctor_get(x_1, 5); +x_15 = lean_ctor_get(x_1, 4); +lean_inc(x_14); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); lean_inc(x_9); -lean_dec(x_3); -x_10 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2; -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -return x_11; +lean_dec(x_1); +x_16 = lean_box(0); +x_17 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_15, x_2, x_16); +x_18 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_11); +lean_ctor_set(x_18, 2, x_12); +lean_ctor_set(x_18, 3, x_13); +lean_ctor_set(x_18, 4, x_17); +lean_ctor_set(x_18, 5, x_14); +lean_ctor_set_uint8(x_18, sizeof(void*)*6, x_10); +return x_18; } } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_3, 1); -lean_inc(x_12); -lean_dec(x_3); -x_13 = lean_box(0); -x_14 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(x_1, x_13, x_12); -return x_14; +lean_dec(x_2); +return x_1; } } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg), 2, 0); -return x_2; +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__4(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(x_1, x_2, x_3); +size_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -return x_4; +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Environment_addExtraName___spec__3(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_2, x_1); -if (x_5 == 0) +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_6; -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_3); -lean_ctor_set(x_6, 1, x_4); -return x_6; -} -else -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_array_uget(x_3, x_2); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_3, x_2, x_8); -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); -lean_dec(x_7); -x_11 = lean_apply_1(x_10, x_4); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; -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 = 1; -x_15 = lean_usize_add(x_2, x_14); -x_16 = lean_array_uset(x_9, x_2, x_12); -x_2 = x_15; -x_3 = x_16; -x_4 = x_13; -goto _start; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } -else -{ -uint8_t x_18; -lean_dec(x_9); -x_18 = !lean_is_exclusive(x_11); -if (x_18 == 0) -{ -return x_11; } -else +LEAN_EXPORT lean_object* l_Lean_Environment_find_x3f(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_11, 0); -x_20 = lean_ctor_get(x_11, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_11); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Lean_SMap_find_x3f_x27___at_Lean_Kernel_Environment_find_x3f___spec__1(x_3, x_2); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Environment_find_x3f___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; -x_3 = lean_st_ref_get(x_2, x_1); -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); -x_6 = lean_array_size(x_4); -x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1(x_6, x_7, x_4, x_5); -return x_8; +lean_object* x_3; +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); +lean_dec(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT uint8_t l_Lean_Environment_contains(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_1); +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); lean_dec(x_1); -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1(x_5, x_6, x_3, x_4); -return x_7; +x_4 = l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(x_3, x_2); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Environment_contains___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_1, x_2, x_3); +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Environment_contains(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Environment_header(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg___boxed), 3, 0); +x_2 = lean_ctor_get(x_1, 5); +lean_inc(x_2); return x_2; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_panic_fn(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_array_get_size(x_2); -x_6 = lean_nat_dec_lt(x_4, x_5); -lean_dec(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -lean_dec(x_3); -x_7 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2; -x_8 = lean_panic_fn(x_2, x_7); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_array_fget(x_2, x_4); -x_10 = lean_box(0); -x_11 = lean_array_fset(x_2, x_4, x_10); -x_12 = lean_apply_1(x_3, x_9); -x_13 = lean_array_fset(x_11, x_4, x_12); -return x_13; -} -} -} -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Environment_header___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg___boxed), 3, 0); +x_2 = l_Lean_Environment_header(x_1); +lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Environment_imports(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_2, x_3); -return x_4; +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 5); +x_3 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Environment_imports___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg___boxed), 3, 0); +x_2 = l_Lean_Environment_imports(x_1); +lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Environment_allImportedModuleNames(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_1, x_2); +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 5); +x_3 = lean_ctor_get(x_2, 3); +lean_inc(x_3); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Environment_allImportedModuleNames___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4___rarg), 2, 0); +x_2 = l_Lean_Environment_allImportedModuleNames(x_1); +lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Environment_setMainModule(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; -x_3 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; -return x_3; -} -} -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__1() { -_start: +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5___boxed), 2, 0); +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 5); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_4, 0); +lean_dec(x_6); +lean_ctor_set(x_4, 0, x_2); return x_1; } -} -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4), 1, 0); +uint32_t 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_uint32(x_4, sizeof(void*)*5); +x_8 = lean_ctor_get(x_4, 1); +x_9 = lean_ctor_get(x_4, 2); +x_10 = lean_ctor_get(x_4, 3); +x_11 = lean_ctor_get(x_4, 4); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_4); +x_12 = lean_alloc_ctor(0, 5, 4); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_8); +lean_ctor_set(x_12, 2, x_9); +lean_ctor_set(x_12, 3, x_10); +lean_ctor_set(x_12, 4, x_11); +lean_ctor_set_uint32(x_12, sizeof(void*)*5, x_7); +lean_ctor_set(x_1, 5, x_12); return x_1; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3), 1, 0); -return x_1; +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint32_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_13 = lean_ctor_get(x_1, 5); +x_14 = lean_ctor_get(x_1, 0); +x_15 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_16 = lean_ctor_get(x_1, 1); +x_17 = lean_ctor_get(x_1, 2); +x_18 = lean_ctor_get(x_1, 3); +x_19 = lean_ctor_get(x_1, 4); +lean_inc(x_13); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_14); +lean_dec(x_1); +x_20 = lean_ctor_get_uint32(x_13, sizeof(void*)*5); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +x_22 = lean_ctor_get(x_13, 2); +lean_inc(x_22); +x_23 = lean_ctor_get(x_13, 3); +lean_inc(x_23); +x_24 = lean_ctor_get(x_13, 4); +lean_inc(x_24); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + lean_ctor_release(x_13, 2); + lean_ctor_release(x_13, 3); + lean_ctor_release(x_13, 4); + x_25 = x_13; +} else { + lean_dec_ref(x_13); + x_25 = lean_box(0); } +if (lean_is_scalar(x_25)) { + x_26 = lean_alloc_ctor(0, 5, 4); +} else { + x_26 = x_25; } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2), 1, 0); -return x_1; +lean_ctor_set(x_26, 0, x_2); +lean_ctor_set(x_26, 1, x_21); +lean_ctor_set(x_26, 2, x_22); +lean_ctor_set(x_26, 3, x_23); +lean_ctor_set(x_26, 4, x_24); +lean_ctor_set_uint32(x_26, sizeof(void*)*5, x_20); +x_27 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_16); +lean_ctor_set(x_27, 2, x_17); +lean_ctor_set(x_27, 3, x_18); +lean_ctor_set(x_27, 4, x_19); +lean_ctor_set(x_27, 5, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*6, x_15); +return x_27; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__5() { +} +LEAN_EXPORT lean_object* l_Lean_Environment_mainModule(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1), 1, 0); -return x_1; +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 5); +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +return x_3; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6() { +LEAN_EXPORT lean_object* l_Lean_Environment_mainModule___boxed(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates), 1, 0); -return x_1; +lean_object* x_2; +x_2 = l_Lean_Environment_mainModule(x_1); +lean_dec(x_1); +return x_2; } } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__7() { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8() { -_start: +if (lean_obj_tag(x_2) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__1; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__2; -x_3 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__4; -x_5 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__5; -x_6 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6; -x_7 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__7; -x_8 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_2); -lean_ctor_set(x_8, 2, x_3); -lean_ctor_set(x_8, 3, x_4); -lean_ctor_set(x_8, 4, x_5); -lean_ctor_set(x_8, 5, x_6); -lean_ctor_set(x_8, 6, x_7); -return x_8; +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_name_eq(x_4, x_1); +if (x_7 == 0) +{ +x_2 = x_6; +goto _start; } -static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8; -return x_1; +lean_object* x_9; +lean_inc(x_5); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_5); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -return x_4; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_3 = lean_ctor_get(x_1, 2); +x_4 = lean_ctor_get(x_3, 1); +x_5 = lean_array_get_size(x_4); +x_6 = l_Lean_Name_hash___override(x_2); +x_7 = 32; +x_8 = lean_uint64_shift_right(x_6, x_7); +x_9 = lean_uint64_xor(x_6, x_8); +x_10 = 16; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = lean_uint64_to_usize(x_12); +x_14 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_15 = 1; +x_16 = lean_usize_sub(x_14, x_15); +x_17 = lean_usize_land(x_13, x_16); +x_18 = lean_array_uget(x_4, x_17); +x_19 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1(x_2, x_18); +lean_dec(x_18); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(x_1, x_2, x_3); +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_getModuleIdxFor_x3f___spec__1(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); -return x_4; +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdxFor_x3f___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5(x_1, x_2); +x_3 = l_Lean_Environment_getModuleIdxFor_x3f(x_1, x_2); lean_dec(x_2); +lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Lean_Environment_isConstructor(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_1, 1); -x_6 = lean_ctor_get(x_1, 2); -x_7 = lean_ctor_get(x_1, 3); -x_8 = lean_ctor_get(x_1, 4); -x_9 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize(x_6, x_2); -if (lean_obj_tag(x_9) == 0) -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_object* x_3; +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); +if (lean_obj_tag(x_3) == 0) { -lean_object* x_11; -x_11 = lean_ctor_get(x_9, 0); -lean_ctor_set(x_1, 2, x_11); -lean_ctor_set(x_9, 0, x_1); -return x_9; +uint8_t x_4; +x_4 = 0; +return x_4; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_9, 0); -x_13 = lean_ctor_get(x_9, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_9); -lean_ctor_set(x_1, 2, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} +lean_object* x_5; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +if (lean_obj_tag(x_5) == 6) +{ +uint8_t x_6; +lean_dec(x_5); +x_6 = 1; +return x_6; } else { -uint8_t x_15; -lean_free_object(x_1); -lean_dec(x_8); -lean_dec(x_7); +uint8_t x_7; lean_dec(x_5); -lean_dec(x_4); -x_15 = !lean_is_exclusive(x_9); -if (x_15 == 0) -{ -return x_9; +x_7 = 0; +return x_7; } -else +} +} +} +LEAN_EXPORT lean_object* l_Lean_Environment_isConstructor___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_9, 0); -x_17 = lean_ctor_get(x_9, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_9); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Environment_isConstructor(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } } +LEAN_EXPORT uint8_t l_Lean_Environment_isSafeDefinition(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = 0; +return x_4; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_19 = lean_ctor_get(x_1, 0); -x_20 = lean_ctor_get(x_1, 1); -x_21 = lean_ctor_get(x_1, 2); -x_22 = lean_ctor_get(x_1, 3); -x_23 = lean_ctor_get(x_1, 4); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_1); -x_24 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize(x_21, x_2); -if (lean_obj_tag(x_24) == 0) +lean_object* x_5; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +if (lean_obj_tag(x_5) == 1) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_27 = x_24; -} else { - lean_dec_ref(x_24); - x_27 = lean_box(0); -} -x_28 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_28, 0, x_19); -lean_ctor_set(x_28, 1, x_20); -lean_ctor_set(x_28, 2, x_25); -lean_ctor_set(x_28, 3, x_22); -lean_ctor_set(x_28, 4, x_23); -if (lean_is_scalar(x_27)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_27; -} -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_26); -return x_29; +lean_object* x_6; uint8_t x_7; lean_object* x_8; +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*)*4); +lean_dec(x_6); +x_8 = lean_box(x_7); +if (lean_obj_tag(x_8) == 1) +{ +uint8_t x_9; +x_9 = 1; +return x_9; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_20); -lean_dec(x_19); -x_30 = lean_ctor_get(x_24, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_24, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_32 = x_24; -} else { - lean_dec_ref(x_24); - x_32 = lean_box(0); +uint8_t x_10; +lean_dec(x_8); +x_10 = 0; +return x_10; } -if (lean_is_scalar(x_32)) { - x_33 = lean_alloc_ctor(1, 2, 0); -} else { - x_33 = x_32; } -lean_ctor_set(x_33, 0, x_30); -lean_ctor_set(x_33, 1, x_31); -return x_33; +else +{ +uint8_t x_11; +lean_dec(x_5); +x_11 = 0; +return x_11; } } } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Environment_isSafeDefinition___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; -x_3 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_EnvExtension_instInhabited(x_1, x_2); +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Environment_isSafeDefinition(x_1, x_2); lean_dec(x_2); -return x_3; +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_Environment_getModuleIdx_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_2); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_2, 2); -x_6 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_5, x_3); -lean_ctor_set(x_2, 2, x_6); -return x_2; -} -else -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_ctor_get(x_2, 2); -x_10 = lean_ctor_get(x_2, 3); -x_11 = lean_ctor_get(x_2, 4); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_2); -x_12 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_9, x_3); -x_13 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_13, 0, x_7); -lean_ctor_set(x_13, 1, x_8); -lean_ctor_set(x_13, 2, x_12); -lean_ctor_set(x_13, 3, x_10); -lean_ctor_set(x_13, 4, x_11); -return x_13; -} +uint8_t x_3; +x_3 = lean_name_eq(x_2, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_setState___rarg___boxed), 3, 0); -return x_2; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_alloc_closure((void*)(l_Lean_Environment_getModuleIdx_x3f___lambda__1___boxed), 2, 1); +lean_closure_set(x_3, 0, x_2); +x_4 = lean_ctor_get(x_1, 5); +x_5 = lean_ctor_get(x_4, 3); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_findIdx_x3f_loop___rarg(x_3, x_5, x_6); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtension_setState___rarg(x_1, x_2, x_3); +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Environment_getModuleIdx_x3f___lambda__1(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); +x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Environment_getModuleIdx_x3f___boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_2); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_2, 2); -x_6 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(x_1, x_5, x_3); -lean_ctor_set(x_2, 2, x_6); -return x_2; +lean_object* x_3; +x_3 = l_Lean_Environment_getModuleIdx_x3f(x_1, x_2); +lean_dec(x_1); +return x_3; } -else -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_ctor_get(x_2, 2); -x_10 = lean_ctor_get(x_2, 3); -x_11 = lean_ctor_get(x_2, 4); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_2); -x_12 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(x_1, x_9, x_3); -x_13 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_13, 0, x_7); -lean_ctor_set(x_13, 1, x_8); -lean_ctor_set(x_13, 2, x_12); -lean_ctor_set(x_13, 3, x_10); -lean_ctor_set(x_13, 4, x_11); -return x_13; } +LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Lean_ConstantInfo_type(x_1); +x_4 = l_Lean_ConstantInfo_levelParams(x_1); +x_5 = l_Lean_Expr_instantiateLevelParams(x_3, x_4, x_2); +lean_dec(x_3); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_modifyState___rarg___boxed), 3, 0); -return x_2; +lean_object* x_3; +x_3 = l_Lean_ConstantInfo_instantiateTypeLevelParams(x_1, x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtension_modifyState___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = 0; +x_4 = l_Lean_ConstantInfo_value_x21(x_1, x_3); +x_5 = l_Lean_ConstantInfo_levelParams(x_1); +x_6 = l_Lean_Expr_instantiateLevelParams(x_4, x_5, x_2); +lean_dec(x_4); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_3, 2); -x_5 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_1, x_2, x_4); -return x_5; +lean_object* x_3; +x_3 = l_Lean_ConstantInfo_instantiateValueLevelParams_x21(x_1, x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_getState___rarg___boxed), 3, 0); +lean_inc(x_2); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_EnvExtension_getState___rarg(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); +x_4 = lean_apply_1(x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_registerEnvExtension___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_3; -x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_1, x_2); +lean_inc(x_3); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_registerEnvExtension(lean_object* x_1) { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__1() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_registerEnvExtension___rarg), 2, 0); +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_mkInitialExtensionStates(lean_object* x_1) { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__2() { _start: { -lean_object* x_2; -x_2 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_2 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__1() { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(10u); -x_2 = lean_unsigned_to_nat(1u); -x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__1___boxed), 2, 0); +return x_1; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__2() { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_mkEmptyEnvironment___lambda__1___closed__1; -x_3 = lean_mk_array(x_2, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__2), 3, 0); +return x_1; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__3() { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_mkEmptyEnvironment___lambda__1___closed__2; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedEnvExtensionInterface___lambda__3___boxed), 4, 0); +return x_1; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__4() { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__6() { _start: { lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +x_1 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__5() { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_mkEmptyEnvironment___lambda__1___closed__4; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_instInhabitedEnvExtensionInterface___closed__3; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__4; +x_3 = l_Lean_instInhabitedEnvExtensionInterface___closed__5; +x_4 = l_Lean_instInhabitedEnvExtensionInterface___closed__2; +x_5 = l_Lean_instInhabitedEnvExtensionInterface___closed__6; +x_6 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_3); +lean_ctor_set(x_6, 4, x_3); +lean_ctor_set(x_6, 5, x_4); +lean_ctor_set(x_6, 6, x_5); +return x_6; } } -static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__6() { +static lean_object* _init_l_Lean_instInhabitedEnvExtensionInterface() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 1; -x_2 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; -x_3 = l_Lean_mkEmptyEnvironment___lambda__1___closed__5; -x_4 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); -return x_4; +lean_object* x_1; +x_1 = l_Lean_instInhabitedEnvExtensionInterface___closed__7; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1(uint32_t x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_3); -if (lean_obj_tag(x_4) == 0) -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_6 = lean_ctor_get(x_4, 0); -x_7 = 0; -x_8 = lean_box(0); -x_9 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_10 = lean_alloc_ctor(0, 5, 5); -lean_ctor_set(x_10, 0, x_8); -lean_ctor_set(x_10, 1, x_9); -lean_ctor_set(x_10, 2, x_9); -lean_ctor_set(x_10, 3, x_9); -lean_ctor_set(x_10, 4, x_9); -lean_ctor_set_uint32(x_10, sizeof(void*)*5, x_1); -lean_ctor_set_uint8(x_10, sizeof(void*)*5 + 4, x_7); -x_11 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; -x_12 = l_Lean_mkEmptyEnvironment___lambda__1___closed__6; -x_13 = l_Lean_NameSet_empty; -x_14 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_14, 0, x_11); -lean_ctor_set(x_14, 1, x_12); -lean_ctor_set(x_14, 2, x_6); -lean_ctor_set(x_14, 3, x_13); -lean_ctor_set(x_14, 4, x_10); -lean_ctor_set(x_4, 0, x_14); -return x_4; +lean_object* x_3; +x_3 = l_Lean_instInhabitedEnvExtensionInterface___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; } -else +} +LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_15 = lean_ctor_get(x_4, 0); -x_16 = lean_ctor_get(x_4, 1); -lean_inc(x_16); -lean_inc(x_15); +lean_object* x_5; +x_5 = l_Lean_instInhabitedEnvExtensionInterface___lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_4); -x_17 = 0; -x_18 = lean_box(0); -x_19 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_20 = lean_alloc_ctor(0, 5, 5); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -lean_ctor_set(x_20, 2, x_19); -lean_ctor_set(x_20, 3, x_19); -lean_ctor_set(x_20, 4, x_19); -lean_ctor_set_uint32(x_20, sizeof(void*)*5, x_1); -lean_ctor_set_uint8(x_20, sizeof(void*)*5 + 4, x_17); -x_21 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; -x_22 = l_Lean_mkEmptyEnvironment___lambda__1___closed__6; -x_23 = l_Lean_NameSet_empty; -x_24 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_24, 0, x_21); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_15); -lean_ctor_set(x_24, 3, x_23); -lean_ctor_set(x_24, 4, x_20); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_16); -return x_25; -} -} -else -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_4); -if (x_26 == 0) -{ -return x_4; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_4, 0); -x_28 = lean_ctor_get(x_4, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_4); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} -} -} -} -static lean_object* _init_l_Lean_mkEmptyEnvironment___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("environment objects cannot be created during initialization", 59, 59); -return x_1; -} -} -static lean_object* _init_l_Lean_mkEmptyEnvironment___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_mkEmptyEnvironment___closed__1; -x_2 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* lean_mk_empty_environment(uint32_t x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_io_initializing(x_2); -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_unbox(x_4); -lean_dec(x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_6); -lean_dec(x_3); -x_7 = lean_box(0); -x_8 = l_Lean_mkEmptyEnvironment___lambda__1(x_1, x_7, x_6); -return x_8; -} -else -{ -uint8_t x_9; -x_9 = !lean_is_exclusive(x_3); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_3, 0); -lean_dec(x_10); -x_11 = l_Lean_mkEmptyEnvironment___closed__2; -lean_ctor_set_tag(x_3, 1); -lean_ctor_set(x_3, 0, x_11); -return x_3; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_3, 1); -lean_inc(x_12); lean_dec(x_3); -x_13 = l_Lean_mkEmptyEnvironment___closed__2; -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -} +lean_dec(x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static uint32_t _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1() { _start: { -uint32_t x_4; lean_object* x_5; -x_4 = lean_unbox_uint32(x_1); -lean_dec(x_1); -x_5 = l_Lean_mkEmptyEnvironment___lambda__1(x_4, x_2, x_3); -lean_dec(x_2); -return x_5; +lean_object* x_1; uint32_t x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_uint32_of_nat(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__2() { _start: { -uint32_t x_3; lean_object* x_4; -x_3 = lean_unbox_uint32(x_1); -lean_dec(x_1); -x_4 = lean_mk_empty_environment(x_3, x_2); +lean_object* x_1; uint32_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1; +x_3 = l_Lean_instToStringImport___closed__2; +x_4 = lean_alloc_ctor(0, 2, 4); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set_uint32(x_4, sizeof(void*)*2, x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_3 = lean_alloc_ctor(0, 2, 0); +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtensionState___rarg), 1, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; -x_4 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__2; -x_5 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1() { _start: { -lean_inc(x_1); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3(lean_object* x_1) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2() { _start: { -lean_object* x_2; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_box(0); +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; return x_2; } } -static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__1() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1833_(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__1___boxed), 3, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__2() { -_start: +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_3 = lean_st_mk_ref(x_2, x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed), 2, 0); -return x_1; -} +return x_3; } -static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed), 1, 0); -return x_1; +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +lean_inc(x_5); +lean_dec(x_3); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; } } -static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__4() { +} +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__4___boxed), 1, 0); +x_1 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef; return x_1; } } -static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__5() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; -x_2 = lean_box(0); -x_3 = l_Lean_instInhabitedPersistentEnvExtension___closed__1; -x_4 = l_Lean_instInhabitedPersistentEnvExtension___closed__2; -x_5 = l_Lean_instInhabitedPersistentEnvExtension___closed__3; -x_6 = l_Lean_instInhabitedPersistentEnvExtension___closed__4; -x_7 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_7, 0, x_1); -lean_ctor_set(x_7, 1, x_2); -lean_ctor_set(x_7, 2, x_3); -lean_ctor_set(x_7, 3, x_4); -lean_ctor_set(x_7, 4, x_5); -lean_ctor_set(x_7, 5, x_6); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; +x_5 = lean_st_ref_get(x_4, x_3); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) { -lean_object* x_5; -x_5 = l_Lean_instInhabitedPersistentEnvExtension___closed__5; +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_array_get_size(x_7); +x_10 = lean_nat_dec_lt(x_1, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_dec(x_7); +lean_dec(x_1); +lean_ctor_set(x_5, 0, x_2); return x_5; } -} -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_instInhabitedPersistentEnvExtension___lambda__1(x_1, x_2, x_3); -lean_dec(x_2); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_free_object(x_5); +x_11 = lean_array_fget(x_7, x_1); +lean_dec(x_7); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_apply_1(x_12, x_8); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_array_push(x_2, x_14); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_1, x_17); lean_dec(x_1); -return x_4; -} +x_1 = x_18; +x_2 = x_16; +x_3 = x_15; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_Lean_instInhabitedPersistentEnvExtension___lambda__2(x_1, x_2); +uint8_t x_20; lean_dec(x_2); lean_dec(x_1); -return x_3; +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) +{ +return x_13; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_13, 0); +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_13); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed(lean_object* x_1) { -_start: +} +} +else { -lean_object* x_2; -x_2 = l_Lean_instInhabitedPersistentEnvExtension___lambda__3(x_1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_5, 0); +x_25 = lean_ctor_get(x_5, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_5); +x_26 = lean_array_get_size(x_24); +x_27 = lean_nat_dec_lt(x_1, x_26); +lean_dec(x_26); +if (x_27 == 0) +{ +lean_object* x_28; +lean_dec(x_24); lean_dec(x_1); -return x_2; +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_2); +lean_ctor_set(x_28, 1, x_25); +return x_28; } +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_array_fget(x_24, x_1); +lean_dec(x_24); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = lean_apply_1(x_30, x_25); +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; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_array_push(x_2, x_32); +x_35 = lean_unsigned_to_nat(1u); +x_36 = lean_nat_add(x_1, x_35); +lean_dec(x_1); +x_1 = x_36; +x_2 = x_34; +x_3 = x_33; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_instInhabitedPersistentEnvExtension___lambda__4(x_1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +lean_dec(x_2); lean_dec(x_1); -return x_2; +x_38 = lean_ctor_get(x_31, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_31, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_40 = x_31; +} else { + lean_dec_ref(x_31); + x_40 = lean_box(0); } +if (lean_is_scalar(x_40)) { + x_41 = lean_alloc_ctor(1, 2, 0); +} else { + x_41 = x_40; } -LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___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_instInhabitedPersistentEnvExtension(x_1, x_2, x_3, x_4); -lean_dec(x_4); -return x_5; +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_5 = l_Lean_instInhabitedPersistentEnvExtensionState___rarg(x_1); -x_6 = lean_ctor_get(x_2, 0); -x_7 = l_Lean_EnvExtension_getState___rarg(x_5, x_6, x_3); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -lean_dec(x_7); -x_9 = l_Lean_instInhabitedModuleData___closed__1; -x_10 = lean_array_get(x_9, x_8, x_4); -lean_dec(x_8); -return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_getModuleEntries___rarg___boxed), 4, 0); +lean_object* x_3; lean_object* x_4; +x_3 = lean_array_get_size(x_1); +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop(x_3, x_1, x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg___closed__1() { _start: { -lean_object* x_5; -x_5 = l_Lean_PersistentEnvExtension_getModuleEntries___rarg(x_1, x_2, x_3, x_4); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid environment extension has been accessed", 47, 47); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg() { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_ctor_get(x_1, 3); -lean_inc(x_4); -lean_dec(x_1); -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_3, 1); -x_7 = lean_apply_2(x_4, x_6, x_2); -lean_ctor_set(x_3, 1, x_7); -return x_3; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_3, 0); -x_9 = lean_ctor_get(x_3, 1); -lean_inc(x_9); -lean_inc(x_8); -lean_dec(x_3); -x_10 = lean_apply_2(x_4, x_9, x_2); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_8); -lean_ctor_set(x_11, 1, x_10); -return x_11; -} +lean_object* x_1; +x_1 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg___closed__1; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_setState___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_addEntry___rarg___lambda__1), 3, 2); -lean_closure_set(x_5, 0, x_1); -lean_closure_set(x_5, 1, x_3); -x_6 = l_Lean_EnvExtension_modifyState___rarg(x_4, x_2, x_5); -lean_dec(x_4); -return x_6; +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1() { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_addEntry___rarg), 3, 0); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Environment", 16, 16); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__2() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l_Lean_instInhabitedPersistentEnvExtensionState___rarg(x_1); -x_5 = lean_ctor_get(x_2, 0); -x_6 = l_Lean_EnvExtension_getState___rarg(x_4, x_5, x_3); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.EnvExtensionInterfaceUnsafe.setState", 41, 41); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__3() { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_getState___rarg___boxed), 3, 0); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__2; +x_3 = lean_unsigned_to_nat(433u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; +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_PersistentEnvExtension_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentEnvExtension_getState___rarg(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___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; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_array_get_size(x_2); +x_6 = lean_nat_dec_lt(x_4, x_5); +lean_dec(x_5); +if (x_6 == 0) { -lean_object* x_4; -x_4 = lean_ctor_get(x_2, 1); -lean_dec(x_4); -lean_ctor_set(x_2, 1, x_1); -return x_2; +lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +x_7 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__3; +x_8 = l_panic___at_Lean_EnvExtensionInterfaceUnsafe_setState___spec__1(x_2, x_7); +return x_8; } else { -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); -lean_dec(x_2); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_1); -return x_6; -} -} +lean_object* x_9; +x_9 = lean_array_fset(x_2, x_4, x_3); +return x_9; } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_setState___rarg___lambda__1), 2, 1); -lean_closure_set(x_5, 0, x_3); -x_6 = l_Lean_EnvExtension_modifyState___rarg(x_4, x_2, x_5); -return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_setState___rarg___boxed), 3, 0); -return x_4; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___boxed), 3, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_PersistentEnvExtension_setState___rarg(x_1, x_2, x_3); +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___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; -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_apply_1(x_1, x_4); -lean_ctor_set(x_2, 1, x_5); -return x_2; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_2, 0); -x_7 = lean_ctor_get(x_2, 1); -lean_inc(x_7); -lean_inc(x_6); -lean_dec(x_2); -x_8 = lean_apply_1(x_1, x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_6); -lean_ctor_set(x_9, 1, x_8); -return x_9; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_modifyState___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_modifyState___rarg___lambda__1), 2, 1); -lean_closure_set(x_5, 0, x_3); -x_6 = l_Lean_EnvExtension_modifyState___rarg(x_4, x_2, x_5); -return x_6; +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__1() { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_modifyState___rarg___boxed), 3, 0); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.EnvExtensionInterfaceUnsafe.modifyState", 44, 44); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2() { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentEnvExtension_modifyState___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__1; +x_3 = lean_unsigned_to_nat(443u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; +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_initFn____x40_Lean_Environment___hyg_2639_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_3 = lean_st_mk_ref(x_2, x_1); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_array_get_size(x_2); +x_6 = lean_nat_dec_lt(x_4, x_5); +lean_dec(x_5); +if (x_6 == 0) { -return x_3; +lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +x_7 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2; +x_8 = l_panic___at_Lean_EnvExtensionInterfaceUnsafe_modifyState___spec__1(x_2, x_7); +return x_8; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_6); -lean_inc(x_5); -lean_dec(x_3); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_5); -lean_ctor_set(x_7, 1, x_6); -return x_7; -} -} +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_array_fget(x_2, x_4); +x_10 = lean_box(0); +x_11 = lean_array_fset(x_2, x_4, x_10); +x_12 = lean_apply_1(x_3, x_9); +x_13 = lean_array_fset(x_11, x_4, x_12); +return x_13; } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__2() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Parser", 6, 6); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___boxed), 3, 0); +return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__3() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Tactic", 6, 6); -return x_1; +lean_object* x_4; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__4() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +x_1 = lean_mk_string_unchecked("Lean.EnvExtensionInterfaceUnsafe.getState", 41, 41); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__5() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___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_Environment___hyg_2684____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2684____closed__4; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__1; +x_3 = lean_unsigned_to_nat(450u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__6() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); -return x_1; -} -} -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__7() { -_start: +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_array_get_size(x_3); +x_6 = lean_nat_dec_lt(x_4, x_5); +lean_dec(x_5); +if (x_6 == 0) { -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_Environment___hyg_2684____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2684____closed__6; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___closed__2; +x_8 = l_panic___rarg(x_1, x_7); +return x_8; } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__8() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("null", 4, 4); -return x_1; -} +lean_object* x_9; +lean_dec(x_1); +x_9 = lean_array_fget(x_3, x_4); +return x_9; } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____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_Environment___hyg_2684____closed__8; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__10() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("exact", 5, 5); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___boxed), 3, 0); +return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__11() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2684____closed__10; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_4; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__12() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Environment___hyg_2684____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_Environment___hyg_2684____closed__13() { -_start: +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; +x_5 = lean_st_ref_get(x_4, x_3); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__12; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__14() { -_start: +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_array_get_size(x_7); +lean_dec(x_7); +lean_ctor_set(x_5, 1, x_1); +lean_ctor_set(x_5, 0, x_9); +x_10 = lean_st_ref_take(x_4, x_8); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +lean_inc(x_5); +x_13 = lean_array_push(x_11, x_5); +x_14 = lean_st_ref_set(x_4, x_13, x_12); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Term", 4, 4); -return x_1; -} +lean_object* x_16; +x_16 = lean_ctor_get(x_14, 0); +lean_dec(x_16); +lean_ctor_set(x_14, 0, x_5); +return x_14; } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__15() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("declName", 8, 8); -return x_1; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_5); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__16() { -_start: +else { -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_Environment___hyg_2684____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____closed__14; -x_4 = l___auto____x40_Lean_Environment___hyg_2684____closed__15; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_19 = lean_ctor_get(x_5, 0); +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_5); +x_21 = lean_array_get_size(x_19); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_1); +x_23 = lean_st_ref_take(x_4, x_20); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +lean_inc(x_22); +x_26 = lean_array_push(x_24, x_22); +x_27 = lean_st_ref_set(x_4, x_26, x_25); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_29 = x_27; +} else { + lean_dec_ref(x_27); + x_29 = lean_box(0); } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__17() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("decl_name%", 10, 10); -return x_1; +if (lean_is_scalar(x_29)) { + x_30 = lean_alloc_ctor(0, 2, 0); +} else { + x_30 = x_29; } +lean_ctor_set(x_30, 0, x_22); +lean_ctor_set(x_30, 1, x_28); +return x_30; } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____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_Environment___hyg_2684____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_Environment___hyg_2684____closed__19() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__18; -x_3 = lean_array_push(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to register environment, extensions can only be registered during initialization", 87, 87); +return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__20() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2() { _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_Environment___hyg_2684____closed__16; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____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); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__1; +x_2 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__21() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__13; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__20; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__22() { -_start: +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_initializing(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_unbox(x_4); +lean_dec(x_4); +if (x_5 == 0) { -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_Environment___hyg_2684____closed__11; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____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); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__23() { -_start: +uint8_t x_6; +lean_dec(x_1); +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__22; -x_3 = lean_array_push(x_1, x_2); +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_3, 0); +lean_dec(x_7); +x_8 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2; +lean_ctor_set_tag(x_3, 1); +lean_ctor_set(x_3, 0, x_8); return x_3; } -} -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__24() { -_start: +else { -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_Environment___hyg_2684____closed__9; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____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); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); +lean_dec(x_3); +x_10 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2; +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__25() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__24; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_dec(x_3); +x_13 = lean_box(0); +x_14 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(x_1, x_13, x_12); +return x_14; } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____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_Environment___hyg_2684____closed__7; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____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); -lean_ctor_set(x_4, 2, x_3); -return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__27() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedModuleData___closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__26; -x_3 = lean_array_push(x_1, x_2); -return x_3; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg), 2, 0); +return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684____closed__28() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Environment___hyg_2684____closed__5; -x_3 = l___auto____x40_Lean_Environment___hyg_2684____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); -lean_ctor_set(x_4, 2, x_3); +lean_object* x_4; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(x_1, x_2, x_3); +lean_dec(x_2); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2684_() { -_start: -{ -lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__28; -return x_1; -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { _start: { uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); +x_5 = lean_usize_dec_lt(x_2, x_1); if (x_5 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -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_ctor_get(x_1, 0); -x_9 = lean_name_eq(x_7, x_8); -lean_dec(x_7); -if (x_9 == 0) -{ -size_t x_10; size_t x_11; -x_10 = 1; -x_11 = lean_usize_add(x_3, x_10); -x_3 = x_11; -goto _start; +lean_object* x_6; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_4); +return x_6; } else { -uint8_t x_13; -x_13 = 1; -return x_13; -} -} -else +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_array_uget(x_3, x_2); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_3, x_2, x_8); +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); +x_11 = lean_apply_1(x_10, x_4); +if (lean_obj_tag(x_11) == 0) { -uint8_t x_14; -x_14 = 0; -return x_14; +lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; +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 = 1; +x_15 = lean_usize_add(x_2, x_14); +x_16 = lean_array_uset(x_9, x_2, x_12); +x_2 = x_15; +x_3 = x_16; +x_4 = x_13; +goto _start; +} +else +{ +uint8_t x_18; +lean_dec(x_9); +x_18 = !lean_is_exclusive(x_11); +if (x_18 == 0) +{ +return x_11; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_11, 0); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg___boxed), 4, 0); -return x_4; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize_loop___closed__1; +x_3 = lean_st_ref_get(x_2, x_1); +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); +x_6 = lean_array_size(x_4); +x_7 = 0; +x_8 = l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1(x_6, x_7, x_4, x_5); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_1); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_2); -return x_5; +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1(x_5, x_6, x_3, x_4); +return x_7; } } -static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1), 2, 0); -return x_1; +lean_object* x_4; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2() { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = l_Lean_persistentEnvExtensionsRef; -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg___boxed), 3, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_ctor_get(x_1, 2); -x_8 = lean_ctor_get(x_1, 3); -x_9 = lean_ctor_get(x_1, 4); -x_10 = lean_ctor_get(x_1, 5); -x_11 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1; -x_12 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_12, 0, x_6); -lean_closure_set(x_12, 1, x_11); -x_13 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_12, x_3); -if (lean_obj_tag(x_13) == 0) +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -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_ctor_set(x_1, 1, x_5); -lean_ctor_set(x_1, 0, x_14); -x_16 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; -x_17 = lean_st_ref_take(x_16, x_15); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -lean_inc(x_1); -x_20 = lean_array_push(x_18, x_1); -x_21 = lean_st_ref_set(x_16, x_20, x_19); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_array_get_size(x_2); +x_6 = lean_nat_dec_lt(x_4, x_5); +lean_dec(x_5); +if (x_6 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_1); -return x_21; +lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +x_7 = l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg___closed__2; +x_8 = lean_panic_fn(x_2, x_7); +return x_8; } 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_1); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_array_fget(x_2, x_4); +x_10 = lean_box(0); +x_11 = lean_array_fset(x_2, x_4, x_10); +x_12 = lean_apply_1(x_3, x_9); +x_13 = lean_array_fset(x_11, x_4, x_12); +return x_13; } } -else +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2(lean_object* x_1) { +_start: { -uint8_t x_26; -lean_free_object(x_1); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -x_26 = !lean_is_exclusive(x_13); -if (x_26 == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -return x_13; +lean_object* x_4; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_2, x_3); +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3(lean_object* x_1) { +_start: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_13, 0); -x_28 = lean_ctor_get(x_13, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_13); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg___boxed), 3, 0); +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_1, x_2); +return x_3; +} } -else +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4(lean_object* x_1) { +_start: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_30 = lean_ctor_get(x_1, 0); -x_31 = lean_ctor_get(x_1, 1); -x_32 = lean_ctor_get(x_1, 2); -x_33 = lean_ctor_get(x_1, 3); -x_34 = lean_ctor_get(x_1, 4); -x_35 = lean_ctor_get(x_1, 5); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_1); -x_36 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1; -x_37 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_37, 0, x_31); -lean_closure_set(x_37, 1, x_36); -x_38 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_37, x_3); -if (lean_obj_tag(x_38) == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_30); -lean_ctor_set(x_41, 2, x_32); -lean_ctor_set(x_41, 3, x_33); -lean_ctor_set(x_41, 4, x_34); -lean_ctor_set(x_41, 5, x_35); -x_42 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; -x_43 = lean_st_ref_take(x_42, x_40); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -lean_inc(x_41); -x_46 = lean_array_push(x_44, x_41); -x_47 = lean_st_ref_set(x_42, x_46, x_45); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_47)) { - lean_ctor_release(x_47, 0); - lean_ctor_release(x_47, 1); - x_49 = x_47; -} else { - lean_dec_ref(x_47); - x_49 = lean_box(0); -} -if (lean_is_scalar(x_49)) { - x_50 = lean_alloc_ctor(0, 2, 0); -} else { - x_50 = x_49; +lean_object* x_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; +return x_3; } -lean_ctor_set(x_50, 0, x_41); -lean_ctor_set(x_50, 1, x_48); -return x_50; } -else +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__1() { +_start: { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_33); -lean_dec(x_32); -lean_dec(x_30); -x_51 = lean_ctor_get(x_38, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_38, 1); -lean_inc(x_52); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_53 = x_38; -} else { - lean_dec_ref(x_38); - x_53 = lean_box(0); -} -if (lean_is_scalar(x_53)) { - x_54 = lean_alloc_ctor(1, 2, 0); -} else { - x_54 = x_53; -} -lean_ctor_set(x_54, 0, x_51); -lean_ctor_set(x_54, 1, x_52); -return x_54; -} -} +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5___boxed), 2, 0); +return x_1; } } -static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid environment extension, '", 32, 32); +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2() { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("' has already been used", 23, 23); +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__4() { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; -x_4 = lean_st_ref_get(x_3, x_2); -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_6 = lean_ctor_get(x_4, 0); -x_7 = lean_ctor_get(x_4, 1); -x_8 = lean_array_get_size(x_6); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_nat_dec_lt(x_9, x_8); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_8); -lean_free_object(x_4); -lean_dec(x_6); -x_11 = lean_box(0); -x_12 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_11, x_7); -return x_12; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2), 1, 0); +return x_1; } -else -{ -size_t x_13; size_t x_14; uint8_t x_15; -x_13 = 0; -x_14 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_15 = l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(x_1, x_6, x_13, x_14); -lean_dec(x_6); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -lean_free_object(x_4); -x_16 = lean_box(0); -x_17 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_16, x_7); -return x_17; } -else +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__5() { +_start: { -lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_18 = lean_ctor_get(x_1, 0); -lean_inc(x_18); -lean_dec(x_1); -x_19 = 1; -x_20 = l_Lean_instToStringImport___closed__1; -x_21 = l_Lean_Name_toString(x_18, x_19, x_20); -x_22 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1; -x_23 = lean_string_append(x_22, x_21); -lean_dec(x_21); -x_24 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2; -x_25 = lean_string_append(x_23, x_24); -x_26 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set_tag(x_4, 1); -lean_ctor_set(x_4, 0, x_26); -return x_4; -} +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1), 1, 0); +return x_1; } } -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_27 = lean_ctor_get(x_4, 0); -x_28 = lean_ctor_get(x_4, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_4); -x_29 = lean_array_get_size(x_27); -x_30 = lean_unsigned_to_nat(0u); -x_31 = lean_nat_dec_lt(x_30, x_29); -if (x_31 == 0) +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6() { +_start: { -lean_object* x_32; lean_object* x_33; -lean_dec(x_29); -lean_dec(x_27); -x_32 = lean_box(0); -x_33 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_32, x_28); -return x_33; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates), 1, 0); +return x_1; } -else -{ -size_t x_34; size_t x_35; uint8_t x_36; -x_34 = 0; -x_35 = lean_usize_of_nat(x_29); -lean_dec(x_29); -x_36 = l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(x_1, x_27, x_34, x_35); -lean_dec(x_27); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; -x_37 = lean_box(0); -x_38 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_37, x_28); -return x_38; } -else +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__7() { +_start: { -lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_39 = lean_ctor_get(x_1, 0); -lean_inc(x_39); -lean_dec(x_1); -x_40 = 1; -x_41 = l_Lean_instToStringImport___closed__1; -x_42 = l_Lean_Name_toString(x_39, x_40, x_41); -x_43 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1; -x_44 = lean_string_append(x_43, x_42); -lean_dec(x_42); -x_45 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2; -x_46 = lean_string_append(x_44, x_45); -x_47 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_47, 0, x_46); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_28); -return x_48; -} -} -} +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8() { _start: { -lean_object* x_5; -x_5 = lean_alloc_closure((void*)(l_Lean_registerPersistentEnvExtensionUnsafe___rarg), 2, 0); -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__1; +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__2; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__4; +x_5 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__5; +x_6 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6; +x_7 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__7; +x_8 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_2); +lean_ctor_set(x_8, 2, x_3); +lean_ctor_set(x_8, 3, x_4); +lean_ctor_set(x_8, 4, x_5); +lean_ctor_set(x_8, 5, x_6); +lean_ctor_set(x_8, 6, x_7); +return x_8; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_EnvExtensionInterfaceUnsafe_imp() { _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_registerPersistentEnvExtensionUnsafe___spec__1___rarg(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_object* x_1; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__8; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_2, x_3); +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__1___rarg(x_1, x_2, x_3); +lean_dec(x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_5; -x_5 = l_Lean_registerPersistentEnvExtensionUnsafe(x_1, x_2, x_3, x_4); -lean_dec(x_4); -return x_5; +lean_object* x_4; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_array_get_size(x_3); -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_nat_dec_lt(x_5, x_4); +lean_object* x_4; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__5(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 3); +lean_inc(x_3); +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_ensureExtensionsArraySize(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); if (x_6 == 0) { -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_2; +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_1, 3); +lean_dec(x_8); +lean_ctor_set(x_1, 3, x_7); +lean_ctor_set(x_4, 0, x_1); +return x_4; } else { -uint8_t x_7; -x_7 = lean_nat_dec_le(x_4, x_4); -if (x_7 == 0) -{ -lean_dec(x_4); -lean_dec(x_3); +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_12 = lean_ctor_get(x_1, 1); +x_13 = lean_ctor_get(x_1, 2); +x_14 = lean_ctor_get(x_1, 4); +x_15 = lean_ctor_get(x_1, 5); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_10); lean_dec(x_1); -return x_2; +x_16 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_13); +lean_ctor_set(x_16, 3, x_9); +lean_ctor_set(x_16, 4, x_14); +lean_ctor_set(x_16, 5, x_15); +lean_ctor_set_uint8(x_16, sizeof(void*)*6, x_11); +lean_ctor_set(x_4, 0, x_16); +return x_4; +} } else { -size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; -x_8 = 0; -x_9 = lean_usize_of_nat(x_4); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_17 = lean_ctor_get(x_4, 0); +x_18 = lean_ctor_get(x_4, 1); +lean_inc(x_18); +lean_inc(x_17); lean_dec(x_4); -x_10 = l_Id_instMonad; -x_11 = l_Array_foldlMUnsafe_fold___rarg(x_10, x_1, x_3, x_8, x_9, x_2); -return x_11; -} +x_19 = lean_ctor_get(x_1, 0); +lean_inc(x_19); +x_20 = lean_ctor_get_uint8(x_1, sizeof(void*)*6); +x_21 = lean_ctor_get(x_1, 1); +lean_inc(x_21); +x_22 = lean_ctor_get(x_1, 2); +lean_inc(x_22); +x_23 = lean_ctor_get(x_1, 4); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 5); +lean_inc(x_24); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + x_25 = x_1; +} else { + lean_dec_ref(x_1); + x_25 = lean_box(0); } +if (lean_is_scalar(x_25)) { + x_26 = lean_alloc_ctor(0, 6, 1); +} else { + x_26 = x_25; } +lean_ctor_set(x_26, 0, x_19); +lean_ctor_set(x_26, 1, x_21); +lean_ctor_set(x_26, 2, x_22); +lean_ctor_set(x_26, 3, x_17); +lean_ctor_set(x_26, 4, x_23); +lean_ctor_set(x_26, 5, x_24); +lean_ctor_set_uint8(x_26, sizeof(void*)*6, x_20); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_18); +return x_27; } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_alloc_closure((void*)(l_Lean_mkStateFromImportedEntries___rarg___lambda__1), 3, 1); -lean_closure_set(x_4, 0, x_1); -x_5 = lean_array_get_size(x_3); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_lt(x_6, x_5); -if (x_7 == 0) -{ -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_2; } else { -uint8_t x_8; -x_8 = lean_nat_dec_le(x_5, x_5); -if (x_8 == 0) +uint8_t x_28; +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_4); +if (x_28 == 0) { -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_2; +return x_4; } else { -size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; -x_9 = 0; -x_10 = lean_usize_of_nat(x_5); -lean_dec(x_5); -x_11 = l_Id_instMonad; -x_12 = l_Array_foldlMUnsafe_fold___rarg(x_11, x_4, x_3, x_9, x_10, x_2); -return x_12; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_4, 0); +x_30 = lean_ctor_get(x_4, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_4); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_mkStateFromImportedEntries___rarg), 3, 0); +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3062_() { -_start: -{ -lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__28; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_apply_1(x_1, x_3); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_2); -lean_ctor_set(x_7, 1, x_6); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_5); -return x_8; +lean_object* x_3; +x_3 = l_Lean_EnvExtension_instInhabited(x_1, x_2); +lean_dec(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; x_4 = !lean_is_exclusive(x_2); if (x_4 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_3); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_3); -lean_ctor_set(x_7, 1, x_5); -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -lean_dec(x_1); -x_9 = lean_apply_2(x_8, x_6, x_3); -lean_ctor_set(x_2, 1, x_9); -lean_ctor_set(x_2, 0, x_7); +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_2, 3); +x_6 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_5, x_3); +lean_ctor_set(x_2, 3, x_6); return x_2; } 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; -x_10 = lean_ctor_get(x_2, 0); -x_11 = lean_ctor_get(x_2, 1); +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; +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*6); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_ctor_get(x_2, 2); +x_11 = lean_ctor_get(x_2, 4); +x_12 = lean_ctor_get(x_2, 5); +x_13 = lean_ctor_get(x_2, 3); +lean_inc(x_12); lean_inc(x_11); +lean_inc(x_13); lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_7); lean_dec(x_2); -lean_inc(x_3); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_3); -lean_ctor_set(x_12, 1, x_10); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_apply_2(x_13, x_11, x_3); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_12); -lean_ctor_set(x_15, 1, x_14); +x_14 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg(x_1, x_13, x_3); +x_15 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_15, 0, x_7); +lean_ctor_set(x_15, 1, x_9); +lean_ctor_set(x_15, 2, x_10); +lean_ctor_set(x_15, 3, x_14); +lean_ctor_set(x_15, 4, x_11); +lean_ctor_set(x_15, 5, x_12); +lean_ctor_set_uint8(x_15, sizeof(void*)*6, x_8); return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState(lean_object* x_1) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_1, 3); -lean_inc(x_3); -lean_dec(x_1); -x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -lean_dec(x_2); -x_5 = l_List_reverse___rarg(x_4); -x_6 = lean_apply_1(x_3, x_5); -return x_6; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_setState___rarg___boxed), 3, 0); +return x_2; } } -static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1() { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("number of local entries: ", 25, 25); -return x_1; +lean_object* x_4; +x_4 = l_Lean_EnvExtension_setState___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2() { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_2, 3); +x_6 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(x_1, x_5, x_3); +lean_ctor_set(x_2, 3, x_6); return x_2; } +else +{ +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; +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*6); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_ctor_get(x_2, 2); +x_11 = lean_ctor_get(x_2, 4); +x_12 = lean_ctor_get(x_2, 5); +x_13 = lean_ctor_get(x_2, 3); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_13); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_7); +lean_dec(x_2); +x_14 = l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(x_1, x_13, x_3); +x_15 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_15, 0, x_7); +lean_ctor_set(x_15, 1, x_9); +lean_ctor_set(x_15, 2, x_10); +lean_ctor_set(x_15, 3, x_14); +lean_ctor_set(x_15, 4, x_11); +lean_ctor_set(x_15, 5, x_12); +lean_ctor_set_uint8(x_15, sizeof(void*)*6, x_8); +return x_15; +} } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(lean_object* x_1) { +} +LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_ctor_get(x_1, 0); -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_List_lengthTRAux___rarg(x_2, x_3); -x_5 = l___private_Init_Data_Repr_0__Nat_reprFast(x_4); -x_6 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_6, 0, x_5); -x_7 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2; -x_8 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; -} -} -static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed), 1, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 2); -lean_inc(x_4); -x_5 = lean_box(0); -x_6 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -lean_inc(x_4); -x_7 = lean_apply_1(x_4, x_6); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_5); -lean_ctor_set(x_8, 1, x_7); -x_9 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_9, 0, x_8); -x_10 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1___boxed), 5, 2); -lean_closure_set(x_10, 0, x_4); -lean_closure_set(x_10, 1, x_5); -lean_inc(x_1); -x_11 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__2), 3, 1); -lean_closure_set(x_11, 0, x_1); -x_12 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__3), 2, 1); -lean_closure_set(x_12, 0, x_1); -x_13 = l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1; -x_14 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_14, 0, x_3); -lean_ctor_set(x_14, 1, x_9); -lean_ctor_set(x_14, 2, x_10); -lean_ctor_set(x_14, 3, x_11); -lean_ctor_set(x_14, 4, x_12); -lean_ctor_set(x_14, 5, x_13); -x_15 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg(x_14, x_2); -return x_15; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_modifyState___rarg___boxed), 3, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg), 2, 0); +x_4 = l_Lean_EnvExtension_modifyState___rarg(x_1, x_2, x_3); +lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; -x_6 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -return x_6; +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_3, 3); +x_5 = l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(x_1, x_2, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(x_1); -lean_dec(x_1); +x_2 = lean_alloc_closure((void*)(l_Lean_EnvExtension_getState___rarg___boxed), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_registerSimplePersistentEnvExtension(x_1, x_2, x_3); +x_4 = l_Lean_EnvExtension_getState___rarg(x_1, x_2, x_3); lean_dec(x_3); +lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_registerEnvExtension___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -x_4 = l_Lean_instInhabitedPersistentEnvExtension(lean_box(0), lean_box(0), lean_box(0), x_3); -lean_dec(x_3); -return x_4; +lean_object* x_3; +x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_registerEnvExtension(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_instInhabited___rarg), 1, 0); -return x_3; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_registerEnvExtension___rarg), 2, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_mkInitialExtensionStates(lean_object* x_1) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = lean_box(0); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_1); -x_6 = l_Lean_PersistentEnvExtension_getState___rarg(x_5, x_2, x_3); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -lean_dec(x_6); -return x_7; +lean_object* x_2; +x_2 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__1() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_getEntries___rarg___boxed), 3, 0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__2() { _start: { -lean_object* x_4; -x_4 = l_Lean_SimplePersistentEnvExtension_getEntries___rarg(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_mkEmptyEnvironment___lambda__1___closed__1; +x_3 = lean_mk_array(x_2, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__3() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = lean_box(0); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_1); -x_6 = l_Lean_PersistentEnvExtension_getState___rarg(x_5, x_2, x_3); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_mkEmptyEnvironment___lambda__1___closed__2; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__4() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_getState___rarg___boxed), 3, 0); -return x_3; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; +x_3 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +x_4 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1(uint32_t x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_3); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_box(0); +x_8 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_9 = lean_alloc_ctor(0, 5, 4); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +lean_ctor_set(x_9, 2, x_8); +lean_ctor_set(x_9, 3, x_8); +lean_ctor_set(x_9, 4, x_8); +lean_ctor_set_uint32(x_9, sizeof(void*)*5, x_1); +x_10 = l_Lean_mkEmptyEnvironment___lambda__1___closed__4; +x_11 = 0; +x_12 = l_Lean_Kernel_instInhabitedDiagnostics___closed__3; +x_13 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; +x_14 = l_Lean_NameSet_empty; +x_15 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_15, 0, x_10); +lean_ctor_set(x_15, 1, x_12); +lean_ctor_set(x_15, 2, x_13); +lean_ctor_set(x_15, 3, x_6); +lean_ctor_set(x_15, 4, x_14); +lean_ctor_set(x_15, 5, x_9); +lean_ctor_set_uint8(x_15, sizeof(void*)*6, x_11); +lean_ctor_set(x_4, 0, x_15); return x_4; } +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_16 = lean_ctor_get(x_4, 0); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_4); +x_18 = lean_box(0); +x_19 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_20 = lean_alloc_ctor(0, 5, 4); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_20, 2, x_19); +lean_ctor_set(x_20, 3, x_19); +lean_ctor_set(x_20, 4, x_19); +lean_ctor_set_uint32(x_20, sizeof(void*)*5, x_1); +x_21 = l_Lean_mkEmptyEnvironment___lambda__1___closed__4; +x_22 = 0; +x_23 = l_Lean_Kernel_instInhabitedDiagnostics___closed__3; +x_24 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; +x_25 = l_Lean_NameSet_empty; +x_26 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_26, 0, x_21); +lean_ctor_set(x_26, 1, x_23); +lean_ctor_set(x_26, 2, x_24); +lean_ctor_set(x_26, 3, x_16); +lean_ctor_set(x_26, 4, x_25); +lean_ctor_set(x_26, 5, x_20); +lean_ctor_set_uint8(x_26, sizeof(void*)*6, x_22); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_17); +return x_27; +} } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_4); +if (x_28 == 0) { -lean_object* x_4; -x_4 = lean_ctor_get(x_2, 1); -lean_dec(x_4); -lean_ctor_set(x_2, 1, x_1); -return x_2; +return x_4; } else { -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); -lean_dec(x_2); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_1); -return x_6; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_4, 0); +x_30 = lean_ctor_get(x_4, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_4); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +static lean_object* _init_l_Lean_mkEmptyEnvironment___closed__1() { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_setState___rarg___lambda__1), 2, 1); -lean_closure_set(x_4, 0, x_3); -x_5 = l_Lean_PersistentEnvExtension_modifyState___rarg(x_1, x_2, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("environment objects cannot be created during initialization", 59, 59); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_mkEmptyEnvironment___closed__2() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_setState___rarg___boxed), 3, 0); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_mkEmptyEnvironment___closed__1; +x_2 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* lean_mk_empty_environment(uint32_t x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_SimplePersistentEnvExtension_setState___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_io_initializing(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_unbox(x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +lean_dec(x_3); +x_7 = lean_box(0); +x_8 = l_Lean_mkEmptyEnvironment___lambda__1(x_1, x_7, x_6); +return x_8; } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) +uint8_t x_9; +x_9 = !lean_is_exclusive(x_3); +if (x_9 == 0) { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_apply_1(x_1, x_4); -lean_ctor_set(x_2, 1, x_5); -return x_2; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_3, 0); +lean_dec(x_10); +x_11 = l_Lean_mkEmptyEnvironment___closed__2; +lean_ctor_set_tag(x_3, 1); +lean_ctor_set(x_3, 0, x_11); +return x_3; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_2, 0); -x_7 = lean_ctor_get(x_2, 1); -lean_inc(x_7); -lean_inc(x_6); -lean_dec(x_2); -x_8 = lean_apply_1(x_1, x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_6); -lean_ctor_set(x_9, 1, x_8); -return x_9; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_dec(x_3); +x_13 = l_Lean_mkEmptyEnvironment___closed__2; +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; -x_4 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_modifyState___rarg___lambda__1), 2, 1); -lean_closure_set(x_4, 0, x_3); -x_5 = l_Lean_PersistentEnvExtension_modifyState___rarg(x_1, x_2, x_4); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_modifyState___rarg___boxed), 3, 0); -return x_3; +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_5 = l_Lean_mkEmptyEnvironment___lambda__1(x_4, x_2, x_3); +lean_dec(x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_SimplePersistentEnvExtension_modifyState___rarg(x_1, x_2, x_3); +uint32_t x_3; lean_object* x_4; +x_3 = lean_unbox_uint32(x_1); lean_dec(x_1); +x_4 = lean_mk_empty_environment(x_3, x_2); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3379_() { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState___rarg(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__28; -return x_1; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Name_quickLt___boxed), 2, 0); -return x_1; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtensionState___rarg), 1, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_7; -x_7 = lean_nat_dec_lt(x_3, x_4); -if (x_7 == 0) -{ -lean_dec(x_3); -return x_2; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_8 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1; -lean_inc(x_3); -x_9 = l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(x_1, x_2, x_8, x_3, x_4, lean_box(0), lean_box(0)); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_nat_dec_le(x_4, x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(x_1, x_11, x_3, x_10, lean_box(0), lean_box(0)); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_add(x_10, x_14); -lean_dec(x_10); -x_2 = x_13; -x_3 = x_15; -x_5 = lean_box(0); -x_6 = lean_box(0); -goto _start; -} -else -{ -lean_dec(x_10); -lean_dec(x_3); -return x_11; -} -} +lean_object* x_4; lean_object* x_5; +x_4 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__2; +x_5 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; -x_3 = lean_box(0); -x_4 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_1, x_2, x_3); -return x_4; +lean_inc(x_1); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_NameSet_empty; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; return x_2; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_2 = lean_array_mk(x_1); -x_3 = lean_array_get_size(x_2); -x_4 = lean_unsigned_to_nat(1u); -x_5 = lean_nat_sub(x_3, x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_3, x_6); -if (x_7 == 0) -{ -uint8_t x_8; -x_8 = lean_nat_dec_le(x_6, x_5); -if (x_8 == 0) -{ -lean_object* x_9; -lean_inc(x_5); -x_9 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(x_3, x_2, x_5, x_5, lean_box(0), lean_box(0)); -lean_dec(x_5); -lean_dec(x_3); -return x_9; -} -else -{ -lean_object* x_10; -x_10 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(x_3, x_2, x_6, x_5, lean_box(0), lean_box(0)); -lean_dec(x_5); -lean_dec(x_3); -return x_10; -} -} -else -{ -lean_dec(x_5); -lean_dec(x_3); +lean_object* x_2; +x_2 = lean_box(0); return x_2; } } -} -static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__1() { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__1___boxed), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__2() { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__2___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__3() { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__3), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__4() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_3 = l_Lean_mkTagDeclarationExtension___closed__1; -x_4 = l_Lean_mkTagDeclarationExtension___closed__2; -x_5 = l_Lean_mkTagDeclarationExtension___closed__3; -x_6 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_6, 0, x_1); -lean_ctor_set(x_6, 1, x_3); -lean_ctor_set(x_6, 2, x_4); -lean_ctor_set(x_6, 3, x_5); -x_7 = l_Lean_registerSimplePersistentEnvExtension___rarg(x_6, x_2); -return x_7; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instInhabitedPersistentEnvExtension___lambda__4___boxed), 1, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___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) { +static lean_object* _init_l_Lean_instInhabitedPersistentEnvExtension___closed__5() { _start: { -lean_object* x_7; -x_7 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_4); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; +x_2 = lean_box(0); +x_3 = l_Lean_instInhabitedPersistentEnvExtension___closed__1; +x_4 = l_Lean_instInhabitedPersistentEnvExtension___closed__2; +x_5 = l_Lean_instInhabitedPersistentEnvExtension___closed__3; +x_6 = l_Lean_instInhabitedPersistentEnvExtension___closed__4; +x_7 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_2); +lean_ctor_set(x_7, 2, x_3); +lean_ctor_set(x_7, 3, x_4); +lean_ctor_set(x_7, 4, x_5); +lean_ctor_set(x_7, 5, x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_2; -x_2 = l_Lean_mkTagDeclarationExtension___lambda__2(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_5; +x_5 = l_Lean_instInhabitedPersistentEnvExtension___closed__5; +return x_5; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_instInhabited___closed__1() { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_NameSet_instInhabited; -x_2 = l_Lean_SimplePersistentEnvExtension_instInhabited___rarg(x_1); -return x_2; +lean_object* x_4; +x_4 = l_Lean_instInhabitedPersistentEnvExtension___lambda__1(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_instInhabited() { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = l_Lean_TagDeclarationExtension_instInhabited___closed__1; -return x_1; +lean_object* x_3; +x_3 = l_Lean_instInhabitedPersistentEnvExtension___lambda__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_TagDeclarationExtension_tag___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = lean_panic_fn(x_1, x_2); -return x_3; +lean_object* x_2; +x_2 = l_Lean_instInhabitedPersistentEnvExtension___lambda__3(x_1); +lean_dec(x_1); +return x_2; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__1() { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4___boxed(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); -return x_1; +lean_object* x_2; +x_2 = l_Lean_instInhabitedPersistentEnvExtension___lambda__4(x_1); +lean_dec(x_1); +return x_2; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__2() { +LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("env.getModuleIdxFor\? declName |>.isNone -- See comment at `TagDeclarationExtension`\n ", 86, 86); -return x_1; +lean_object* x_5; +x_5 = l_Lean_instInhabitedPersistentEnvExtension(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__3() { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_TagDeclarationExtension_tag___closed__1; -x_2 = l_Lean_TagDeclarationExtension_tag___closed__2; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_5 = l_Lean_instInhabitedPersistentEnvExtensionState___rarg(x_1); +x_6 = lean_ctor_get(x_2, 0); +x_7 = l_Lean_EnvExtension_getState___rarg(x_5, x_6, x_3); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +x_9 = l_Lean_instInhabitedModuleData___closed__1; +x_10 = lean_array_get(x_9, x_8, x_4); +lean_dec(x_8); +return x_10; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__4() { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.TagDeclarationExtension.tag", 32, 32); -return x_1; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_getModuleEntries___rarg___boxed), 4, 0); +return x_4; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__5() { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; -x_2 = l_Lean_TagDeclarationExtension_tag___closed__4; -x_3 = lean_unsigned_to_nat(628u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_TagDeclarationExtension_tag___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_5; +x_5 = l_Lean_PersistentEnvExtension_getModuleEntries___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_tag(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; -x_4 = l_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); -if (lean_obj_tag(x_4) == 0) +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 3); +lean_inc(x_4); +lean_dec(x_1); +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -lean_object* x_5; -x_5 = l_Lean_PersistentEnvExtension_addEntry___rarg(x_1, x_2, x_3); -return x_5; +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 1); +x_7 = lean_apply_2(x_4, x_6, x_2); +lean_ctor_set(x_3, 1, x_7); +return x_3; } else { -lean_object* x_6; lean_object* x_7; -lean_dec(x_4); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); +lean_inc(x_8); lean_dec(x_3); -lean_dec(x_1); -x_6 = l_Lean_TagDeclarationExtension_tag___closed__5; -x_7 = l_panic___at_Lean_TagDeclarationExtension_tag___spec__1(x_2, x_6); -return x_7; +x_10 = lean_apply_2(x_4, x_9, x_2); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; } } } -LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_6 = lean_nat_add(x_3, x_4); -x_7 = lean_unsigned_to_nat(2u); -x_8 = lean_nat_div(x_6, x_7); -lean_dec(x_6); -x_9 = lean_array_fget(x_1, x_8); -x_10 = l_Lean_Name_quickLt(x_9, x_2); -if (x_10 == 0) -{ -uint8_t x_11; +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_addEntry___rarg___lambda__1), 3, 2); +lean_closure_set(x_5, 0, x_1); +lean_closure_set(x_5, 1, x_3); +x_6 = l_Lean_EnvExtension_modifyState___rarg(x_4, x_2, x_5); lean_dec(x_4); -x_11 = l_Lean_Name_quickLt(x_2, x_9); -lean_dec(x_9); -if (x_11 == 0) -{ -uint8_t x_12; -lean_dec(x_8); -lean_dec(x_3); -x_12 = 1; -return x_12; +return x_6; } -else -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_eq(x_8, x_13); -if (x_14 == 0) +} +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_8, x_15); -lean_dec(x_8); -x_17 = lean_nat_dec_lt(x_16, x_3); -if (x_17 == 0) +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_addEntry___rarg), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -x_4 = x_16; -x_5 = lean_box(0); -goto _start; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = l_Lean_instInhabitedPersistentEnvExtensionState___rarg(x_1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = l_Lean_EnvExtension_getState___rarg(x_4, x_5, x_3); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +return x_7; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_19; -lean_dec(x_16); -lean_dec(x_3); -x_19 = 0; -return x_19; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_getState___rarg___boxed), 3, 0); +return x_4; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_20; -lean_dec(x_8); +lean_object* x_4; +x_4 = l_Lean_PersistentEnvExtension_getState___rarg(x_1, x_2, x_3); lean_dec(x_3); -x_20 = 0; -return x_20; -} +lean_dec(x_2); +return x_4; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -lean_dec(x_9); -lean_dec(x_3); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_add(x_8, x_21); -lean_dec(x_8); -x_23 = lean_nat_dec_le(x_22, x_4); -if (x_23 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -uint8_t x_24; -lean_dec(x_22); +lean_object* x_4; +x_4 = lean_ctor_get(x_2, 1); lean_dec(x_4); -x_24 = 0; -return x_24; +lean_ctor_set(x_2, 1, x_1); +return x_2; } else { -x_3 = x_22; -x_5 = lean_box(0); -goto _start; +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec(x_2); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_1); +return x_6; +} } } +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_setState___rarg___lambda__1), 2, 1); +lean_closure_set(x_5, 0, x_3); +x_6 = l_Lean_EnvExtension_modifyState___rarg(x_4, x_2, x_5); +return x_6; } } -static lean_object* _init_l_Lean_TagDeclarationExtension_isTagged___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_NameSet_instInhabited; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_setState___rarg___boxed), 3, 0); +return x_4; } } -LEAN_EXPORT uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_NameSet_instInhabited; -x_6 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_5, x_1, x_2); -x_7 = l_Lean_NameSet_contains(x_6, x_3); -lean_dec(x_6); -return x_7; +x_4 = l_Lean_PersistentEnvExtension_setState___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } -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; uint8_t x_15; -x_8 = lean_ctor_get(x_4, 0); -lean_inc(x_8); -lean_dec(x_4); -x_9 = l_Lean_TagDeclarationExtension_isTagged___closed__1; -x_10 = l_Lean_PersistentEnvExtension_getModuleEntries___rarg(x_9, x_1, x_2, x_8); -lean_dec(x_8); -x_11 = lean_array_get_size(x_10); -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_sub(x_11, x_12); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_11); -lean_dec(x_11); -if (x_15 == 0) -{ -uint8_t x_16; -lean_dec(x_13); -lean_dec(x_10); -x_16 = 0; -return x_16; } -else +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_17; -x_17 = lean_nat_dec_le(x_14, x_13); -if (x_17 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -uint8_t x_18; -lean_dec(x_13); -lean_dec(x_10); -x_18 = 0; -return x_18; +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_apply_1(x_1, x_4); +lean_ctor_set(x_2, 1, x_5); +return x_2; } else { -uint8_t x_19; -x_19 = l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1(x_10, x_3, x_14, x_13, lean_box(0)); -lean_dec(x_10); -return x_19; -} -} +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_2); +x_8 = lean_apply_1(x_1, x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_6); +lean_ctor_set(x_9, 1, x_8); +return x_9; } } } -LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_6; lean_object* x_7; -x_6 = l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_2); -lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_modifyState___rarg___lambda__1), 2, 1); +lean_closure_set(x_5, 0, x_3); +x_6 = l_Lean_EnvExtension_modifyState___rarg(x_4, x_2, x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_isTagged___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_TagDeclarationExtension_isTagged(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_5 = lean_box(x_4); -return x_5; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentEnvExtension_modifyState___rarg___boxed), 3, 0); +return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3542_() { +LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__28; -return x_1; +lean_object* x_4; +x_4 = l_Lean_PersistentEnvExtension_modifyState___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_3055_(lean_object* x_1) { _start: { -if (lean_obj_tag(x_2) == 0) +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_3 = lean_st_mk_ref(x_2, x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) { -return x_1; +return x_3; } else { -lean_object* x_3; 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 = lean_ctor_get(x_2, 0); -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_ctor_get(x_2, 3); -x_7 = l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(x_1, x_3); +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_4); -lean_ctor_set(x_8, 1, x_5); -x_9 = lean_array_push(x_7, x_8); -x_1 = x_9; -x_2 = x_6; -goto _start; +lean_dec(x_3); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__1() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg___boxed), 2, 0); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__2() { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_3 = l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(x_2, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__3() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed), 1, 0); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__4() { _start: { -lean_object* x_5; -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__5() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = lean_ctor_get(x_2, 0); -lean_inc(x_3); -x_4 = lean_ctor_get(x_2, 1); -lean_inc(x_4); -lean_dec(x_2); -x_5 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_1, x_3, x_4); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_3100____closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_mkMapDeclarationExtension___rarg___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_mkMapDeclarationExtension___rarg___closed__2() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkMapDeclarationExtension___rarg___lambda__2), 2, 0); +x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__7() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_3 = lean_box(0); -x_4 = lean_alloc_closure((void*)(l_Lean_mkMapDeclarationExtension___rarg___lambda__1___boxed), 4, 1); -lean_closure_set(x_4, 0, x_3); -x_5 = l_Lean_mkMapDeclarationExtension___rarg___closed__1; -x_6 = l_Lean_mkMapDeclarationExtension___rarg___closed__2; -x_7 = lean_alloc_closure((void*)(l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed), 1, 0); -x_8 = l_Lean_instInhabitedPersistentEnvExtension___closed__4; -x_9 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_5); -lean_ctor_set(x_9, 2, x_4); -lean_ctor_set(x_9, 3, x_6); -lean_ctor_set(x_9, 4, x_7); -lean_ctor_set(x_9, 5, x_8); -x_10 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg(x_9, x_2); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_3100____closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__8() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_mkMapDeclarationExtension___rarg), 2, 0); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__9() { _start: { -lean_object* x_3; -x_3 = l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(x_1, x_2); -lean_dec(x_2); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__10() { _start: { -lean_object* x_2; -x_2 = l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("exact", 5, 5); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__11() { _start: { -lean_object* x_5; -x_5 = l_Lean_mkMapDeclarationExtension___rarg___lambda__1(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_3100____closed__10; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_MapDeclarationExtension_instInhabited___closed__1() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__12() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = l_Lean_instInhabitedPersistentEnvExtension(lean_box(0), lean_box(0), lean_box(0), x_1); -return x_2; -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l___auto____x40_Lean_Environment___hyg_3100____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; } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_instInhabited(lean_object* x_1) { +} +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__13() { _start: { -lean_object* x_2; -x_2 = l_Lean_MapDeclarationExtension_instInhabited___closed__1; -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instInhabitedModuleData___closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__12; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__14() { _start: { -lean_object* x_3; -x_3 = lean_panic_fn(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; } } -static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__1() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__15() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("env.getModuleIdxFor\? declName |>.isNone -- See comment at `MapDeclarationExtension`\n ", 86, 86); +x_1 = lean_mk_string_unchecked("declName", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__2() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__16() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_TagDeclarationExtension_tag___closed__1; -x_2 = l_Lean_MapDeclarationExtension_insert___rarg___closed__1; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____closed__14; +x_4 = l___auto____x40_Lean_Environment___hyg_3100____closed__15; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__3() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__17() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.MapDeclarationExtension.insert", 35, 35); +x_1 = lean_mk_string_unchecked("decl_name%", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__4() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__18() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; -x_2 = l_Lean_MapDeclarationExtension_insert___rarg___closed__3; -x_3 = lean_unsigned_to_nat(659u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_MapDeclarationExtension_insert___rarg___closed__2; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l___auto____x40_Lean_Environment___hyg_3100____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; } } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__19() { _start: { -lean_object* x_5; -x_5 = l_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_3); -lean_ctor_set(x_6, 1, x_4); -x_7 = l_Lean_PersistentEnvExtension_addEntry___rarg(x_1, x_2, x_6); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instInhabitedModuleData___closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__18; +x_3 = lean_array_push(x_1, x_2); +return x_3; } -else -{ -lean_object* x_8; lean_object* x_9; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_8 = l_Lean_MapDeclarationExtension_insert___rarg___closed__4; -x_9 = l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(x_2, x_8); -return x_9; } +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____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_Environment___hyg_3100____closed__16; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____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); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__21() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_insert___rarg), 4, 0); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__13; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__20; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__22() { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_6 = lean_nat_add(x_3, x_4); -x_7 = lean_unsigned_to_nat(2u); -x_8 = lean_nat_div(x_6, x_7); -lean_dec(x_6); -x_9 = lean_array_fget(x_1, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_2, 0); -x_12 = l_Lean_Name_quickLt(x_10, x_11); -if (x_12 == 0) -{ -uint8_t x_13; -lean_dec(x_4); -x_13 = l_Lean_Name_quickLt(x_11, x_10); -lean_dec(x_10); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_8); -lean_dec(x_3); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_9); -return x_14; +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_Environment___hyg_3100____closed__11; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____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); +lean_ctor_set(x_4, 2, x_3); +return x_4; } -else -{ -lean_object* x_15; uint8_t x_16; -lean_dec(x_9); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_nat_dec_eq(x_8, x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_8, x_17); -lean_dec(x_8); -x_19 = lean_nat_dec_lt(x_18, x_3); -if (x_19 == 0) -{ -x_4 = x_18; -x_5 = lean_box(0); -goto _start; } -else +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__23() { +_start: { -lean_object* x_21; -lean_dec(x_18); -lean_dec(x_3); -x_21 = lean_box(0); -return x_21; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instInhabitedModuleData___closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__22; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__24() { +_start: { -lean_object* x_22; -lean_dec(x_8); -lean_dec(x_3); -x_22 = lean_box(0); -return x_22; -} +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_Environment___hyg_3100____closed__9; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____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); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -else +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__25() { +_start: { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_add(x_8, x_23); -lean_dec(x_8); -x_25 = lean_nat_dec_le(x_24, x_4); -if (x_25 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instInhabitedModuleData___closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__24; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__26() { +_start: { -lean_object* x_26; -lean_dec(x_24); -lean_dec(x_4); -x_26 = lean_box(0); -return x_26; +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_Environment___hyg_3100____closed__7; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____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); +lean_ctor_set(x_4, 2, x_3); +return x_4; } -else +} +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____closed__27() { +_start: { -x_3 = x_24; -x_5 = lean_box(0); -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instInhabitedModuleData___closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_3100____closed__26; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100____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_Environment___hyg_3100____closed__5; +x_3 = l___auto____x40_Lean_Environment___hyg_3100____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); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1(lean_object* x_1) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3100_() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg___boxed), 5, 0); -return x_2; +lean_object* x_1; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__28; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { _start: { -lean_object* x_5; -x_5 = l_Lean_Environment_getModuleIdxFor_x3f(x_3, x_4); -if (lean_obj_tag(x_5) == 0) +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; lean_object* x_8; -lean_dec(x_1); -x_6 = lean_box(0); -x_7 = l_Lean_PersistentEnvExtension_getState___rarg(x_6, x_2, x_3); -x_8 = l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(x_7, x_4); -lean_dec(x_4); +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +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_ctor_get(x_1, 0); +x_9 = lean_name_eq(x_7, x_8); lean_dec(x_7); -return x_8; +if (x_9 == 0) +{ +size_t x_10; size_t x_11; +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_3 = x_11; +goto _start; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_9 = lean_ctor_get(x_5, 0); -lean_inc(x_9); -lean_dec(x_5); -x_10 = lean_box(0); -x_11 = l_Lean_PersistentEnvExtension_getModuleEntries___rarg(x_10, x_2, x_3, x_9); -lean_dec(x_9); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_4); -lean_ctor_set(x_12, 1, x_1); -x_13 = lean_array_get_size(x_11); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_13, x_14); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_lt(x_16, x_13); -lean_dec(x_13); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_12); -lean_dec(x_11); -x_18 = lean_box(0); -return x_18; -} -else -{ -uint8_t x_19; -x_19 = lean_nat_dec_le(x_16, x_15); -if (x_19 == 0) -{ -lean_object* x_20; -lean_dec(x_15); -lean_dec(x_12); -lean_dec(x_11); -x_20 = lean_box(0); -return x_20; -} -else -{ -lean_object* x_21; -x_21 = l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg(x_11, x_12, x_16, x_15, lean_box(0)); -lean_dec(x_12); -lean_dec(x_11); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; -x_22 = lean_box(0); -return x_22; +uint8_t x_13; +x_13 = 1; +return x_13; } -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_21); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 0); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -lean_ctor_set(x_21, 0, x_25); -return x_21; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_21, 0); -lean_inc(x_26); -lean_dec(x_21); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -return x_28; -} -} -} -} +uint8_t x_14; +x_14 = 0; +return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_find_x3f___rarg___boxed), 4, 0); -return x_2; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg___boxed), 4, 0); +return x_4; } } -LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_6; -x_6 = l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_6; +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_1); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1() { _start: { -lean_object* x_5; -x_5 = l_Lean_MapDeclarationExtension_find_x3f___rarg(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1), 2, 0); +return x_1; } } -LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2() { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_6 = lean_nat_add(x_3, x_4); -x_7 = lean_unsigned_to_nat(2u); -x_8 = lean_nat_div(x_6, x_7); -lean_dec(x_6); -x_9 = lean_array_fget(x_1, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_ctor_get(x_2, 0); -x_12 = l_Lean_Name_quickLt(x_10, x_11); -if (x_12 == 0) -{ -uint8_t x_13; -lean_dec(x_4); -x_13 = l_Lean_Name_quickLt(x_11, x_10); -lean_dec(x_10); -if (x_13 == 0) -{ -uint8_t x_14; -lean_dec(x_8); -lean_dec(x_3); -x_14 = 1; -return x_14; +lean_object* x_1; +x_1 = l_Lean_persistentEnvExtensionsRef; +return x_1; } -else +} +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_15; uint8_t x_16; -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_nat_dec_eq(x_8, x_15); -if (x_16 == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_8, x_17); -lean_dec(x_8); -x_19 = lean_nat_dec_lt(x_18, x_3); -if (x_19 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_1, 2); +x_8 = lean_ctor_get(x_1, 3); +x_9 = lean_ctor_get(x_1, 4); +x_10 = lean_ctor_get(x_1, 5); +x_11 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1; +x_12 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); +lean_closure_set(x_12, 0, x_6); +lean_closure_set(x_12, 1, x_11); +x_13 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_12, x_3); +if (lean_obj_tag(x_13) == 0) { -x_4 = x_18; -x_5 = lean_box(0); -goto _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; lean_object* x_20; lean_object* x_21; uint8_t x_22; +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_ctor_set(x_1, 1, x_5); +lean_ctor_set(x_1, 0, x_14); +x_16 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; +x_17 = lean_st_ref_take(x_16, x_15); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +lean_inc(x_1); +x_20 = lean_array_push(x_18, x_1); +x_21 = lean_st_ref_set(x_16, x_20, x_19); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -uint8_t x_21; -lean_dec(x_18); -lean_dec(x_3); -x_21 = 0; +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_1); return x_21; } -} else { -uint8_t x_22; -lean_dec(x_8); -lean_dec(x_3); -x_22 = 0; -return x_22; -} +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_1); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; +uint8_t x_26; +lean_free_object(x_1); lean_dec(x_10); -lean_dec(x_3); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_add(x_8, x_23); +lean_dec(x_9); lean_dec(x_8); -x_25 = lean_nat_dec_le(x_24, x_4); -if (x_25 == 0) +lean_dec(x_7); +lean_dec(x_5); +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) { -uint8_t x_26; -lean_dec(x_24); -lean_dec(x_4); -x_26 = 0; -return x_26; +return x_13; } else { -x_3 = x_24; -x_5 = lean_box(0); -goto _start; -} +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_13, 0); +x_28 = lean_ctor_get(x_13, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_13); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } -LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg___boxed), 5, 0); -return x_2; -} -} -LEAN_EXPORT uint8_t l_Lean_MapDeclarationExtension_contains___rarg(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_Environment_getModuleIdxFor_x3f(x_3, x_4); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_30 = lean_ctor_get(x_1, 0); +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_ctor_get(x_1, 2); +x_33 = lean_ctor_get(x_1, 3); +x_34 = lean_ctor_get(x_1, 4); +x_35 = lean_ctor_get(x_1, 5); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); lean_dec(x_1); -x_6 = lean_box(0); -x_7 = l_Lean_PersistentEnvExtension_getState___rarg(x_6, x_2, x_3); -x_8 = l_Lean_NameMap_contains___rarg(x_7, x_4); -lean_dec(x_4); -lean_dec(x_7); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_9 = lean_ctor_get(x_5, 0); -lean_inc(x_9); -lean_dec(x_5); -x_10 = lean_box(0); -x_11 = l_Lean_PersistentEnvExtension_getModuleEntries___rarg(x_10, x_2, x_3, x_9); -lean_dec(x_9); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_4); -lean_ctor_set(x_12, 1, x_1); -x_13 = lean_array_get_size(x_11); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_13, x_14); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_lt(x_16, x_13); -lean_dec(x_13); -if (x_17 == 0) -{ -uint8_t x_18; -lean_dec(x_15); -lean_dec(x_12); -lean_dec(x_11); -x_18 = 0; -return x_18; -} -else -{ -uint8_t x_19; -x_19 = lean_nat_dec_le(x_16, x_15); -if (x_19 == 0) -{ -uint8_t x_20; -lean_dec(x_15); -lean_dec(x_12); -lean_dec(x_11); -x_20 = 0; -return x_20; -} -else +x_36 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1; +x_37 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); +lean_closure_set(x_37, 0, x_31); +lean_closure_set(x_37, 1, x_36); +x_38 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_37, x_3); +if (lean_obj_tag(x_38) == 0) { -uint8_t x_21; -x_21 = l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(x_11, x_12, x_16, x_15, lean_box(0)); -lean_dec(x_12); -lean_dec(x_11); -return x_21; -} -} +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_30); +lean_ctor_set(x_41, 2, x_32); +lean_ctor_set(x_41, 3, x_33); +lean_ctor_set(x_41, 4, x_34); +lean_ctor_set(x_41, 5, x_35); +x_42 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; +x_43 = lean_st_ref_take(x_42, x_40); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +lean_inc(x_41); +x_46 = lean_array_push(x_44, x_41); +x_47 = lean_st_ref_set(x_42, x_46, x_45); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); } +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(0, 2, 0); +} else { + x_50 = x_49; } +lean_ctor_set(x_50, 0, x_41); +lean_ctor_set(x_50, 1, x_48); +return x_50; } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_contains___rarg___boxed), 4, 0); -return x_2; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_30); +x_51 = lean_ctor_get(x_38, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_38, 1); +lean_inc(x_52); +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + x_53 = x_38; +} else { + lean_dec_ref(x_38); + x_53 = lean_box(0); } +if (lean_is_scalar(x_53)) { + x_54 = lean_alloc_ctor(1, 2, 0); +} else { + x_54 = x_53; } -LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -uint8_t x_6; lean_object* x_7; -x_6 = l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_2); -lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; +lean_ctor_set(x_54, 0, x_51); +lean_ctor_set(x_54, 1, x_52); +return x_54; } } -LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; lean_object* x_6; -x_5 = l_Lean_MapDeclarationExtension_contains___rarg(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -x_6 = lean_box(x_5); -return x_6; } } -LEAN_EXPORT lean_object* l_Lean_saveModuleData___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1() { _start: { -lean_object* x_5; -x_5 = lean_save_module_data(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid environment extension, '", 32, 32); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_readModuleData___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2() { _start: { -lean_object* x_3; -x_3 = lean_read_module_data(x_1, x_2); -lean_dec(x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' has already been used", 23, 23); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_2, x_3); -if (x_6 == 0) +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; +x_4 = lean_st_ref_get(x_3, x_2); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) { -lean_object* x_7; size_t x_8; lean_object* x_9; -lean_dec(x_4); -x_7 = lean_array_uget(x_1, x_2); -x_8 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_9 = lean_compacted_region_free(x_8, x_5); -if (lean_obj_tag(x_9) == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_array_get_size(x_6); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_lt(x_9, x_8); +if (x_10 == 0) { -lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = 1; -x_13 = lean_usize_add(x_2, x_12); -x_2 = x_13; -x_4 = x_10; -x_5 = x_11; -goto _start; +lean_object* x_11; lean_object* x_12; +lean_dec(x_8); +lean_free_object(x_4); +lean_dec(x_6); +x_11 = lean_box(0); +x_12 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_11, x_7); +return x_12; } else { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_9); +size_t x_13; size_t x_14; uint8_t x_15; +x_13 = 0; +x_14 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_15 = l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(x_1, x_6, x_13, x_14); +lean_dec(x_6); if (x_15 == 0) { -return x_9; +lean_object* x_16; lean_object* x_17; +lean_free_object(x_4); +x_16 = lean_box(0); +x_17 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_16, x_7); +return x_17; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_9, 0); -x_17 = lean_ctor_get(x_9, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_9); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_18 = lean_ctor_get(x_1, 0); +lean_inc(x_18); +lean_dec(x_1); +x_19 = 1; +x_20 = l_Lean_instToStringImport___closed__1; +x_21 = l_Lean_Name_toString(x_18, x_19, x_20); +x_22 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1; +x_23 = lean_string_append(x_22, x_21); +lean_dec(x_21); +x_24 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2; +x_25 = lean_string_append(x_23, x_24); +x_26 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set_tag(x_4, 1); +lean_ctor_set(x_4, 0, x_26); +return x_4; } } } else { -lean_object* x_19; -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_4); -lean_ctor_set(x_19, 1, x_5); -return x_19; -} -} -} -LEAN_EXPORT lean_object* lean_environment_free_regions(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_3 = lean_ctor_get(x_1, 4); -lean_inc(x_3); -lean_dec(x_1); -x_4 = lean_ctor_get(x_3, 2); -lean_inc(x_4); -lean_dec(x_3); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_lt(x_6, x_5); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -lean_dec(x_5); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_27 = lean_ctor_get(x_4, 0); +x_28 = lean_ctor_get(x_4, 1); +lean_inc(x_28); +lean_inc(x_27); lean_dec(x_4); -x_8 = lean_box(0); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_2); -return x_9; +x_29 = lean_array_get_size(x_27); +x_30 = lean_unsigned_to_nat(0u); +x_31 = lean_nat_dec_lt(x_30, x_29); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_29); +lean_dec(x_27); +x_32 = lean_box(0); +x_33 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_32, x_28); +return x_33; } else { -uint8_t x_10; -x_10 = lean_nat_dec_le(x_5, x_5); -if (x_10 == 0) +size_t x_34; size_t x_35; uint8_t x_36; +x_34 = 0; +x_35 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_36 = l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(x_1, x_27, x_34, x_35); +lean_dec(x_27); +if (x_36 == 0) { -lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -lean_dec(x_4); -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -return x_12; +lean_object* x_37; lean_object* x_38; +x_37 = lean_box(0); +x_38 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_37, x_28); +return x_38; } else { -size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; -x_13 = 0; -x_14 = lean_usize_of_nat(x_5); -lean_dec(x_5); -x_15 = lean_box(0); -x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1(x_4, x_13, x_14, x_15, x_2); -lean_dec(x_4); -return x_16; +lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_39 = lean_ctor_get(x_1, 0); +lean_inc(x_39); +lean_dec(x_1); +x_40 = 1; +x_41 = l_Lean_instToStringImport___closed__1; +x_42 = l_Lean_Name_toString(x_39, x_40, x_41); +x_43 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1; +x_44 = lean_string_append(x_43, x_42); +lean_dec(x_42); +x_45 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2; +x_46 = lean_string_append(x_44, x_45); +x_47 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_47, 0, x_46); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_28); +return x_48; } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Lean_registerPersistentEnvExtensionUnsafe___rarg), 2, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg___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_8 = l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1(x_1, x_6, x_7, x_4, x_5); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1___rarg(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_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) -{ +lean_object* x_4; +x_4 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(x_1, x_2, x_3); +lean_dec(x_2); return x_4; } -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; -x_6 = lean_array_uget(x_4, x_3); -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_array_uset(x_4, x_3, x_7); -x_9 = l_Lean_instInhabitedEnvExtensionState; -x_10 = l_Lean_PersistentEnvExtension_getState___rarg(x_9, x_6, x_1); -x_11 = lean_ctor_get(x_6, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_6, 4); -lean_inc(x_12); -lean_dec(x_6); -x_13 = lean_apply_1(x_12, x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_11); -lean_ctor_set(x_14, 1, x_13); -x_15 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_17 = lean_array_uset(x_8, x_3, x_14); -x_3 = x_16; -x_4 = x_17; -goto _start; } +LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___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_registerPersistentEnvExtensionUnsafe(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_array_get_size(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_4); if (x_6 == 0) { -lean_object* x_7; size_t x_8; size_t x_9; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -switch (lean_obj_tag(x_7)) { -case 0: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_7, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_7, 1); -lean_inc(x_11); -lean_dec(x_7); -lean_inc(x_1); -x_12 = lean_apply_3(x_1, x_5, x_10, x_11); -x_3 = x_9; -x_5 = x_12; -goto _start; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_2; } -case 1: +else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_7, 0); -lean_inc(x_14); -lean_dec(x_7); -lean_inc(x_1); -x_15 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_1, x_14, x_5); -lean_dec(x_14); -x_3 = x_9; -x_5 = x_15; -goto _start; -} -default: +uint8_t x_7; +x_7 = lean_nat_dec_le(x_4, x_4); +if (x_7 == 0) { -x_3 = x_9; -goto _start; -} -} +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_2; } else { -lean_dec(x_1); -return x_5; -} +size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = 0; +x_9 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_10 = l_Id_instMonad; +x_11 = l_Array_foldlMUnsafe_fold___rarg(x_10, x_1, x_3, x_8, x_9, x_2); +return x_11; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg___boxed), 5, 0); -return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_5); -lean_dec(x_1); -return x_6; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -lean_inc(x_1); -x_11 = lean_apply_3(x_1, x_6, x_9, x_10); -x_12 = lean_unsigned_to_nat(1u); -x_13 = lean_nat_add(x_5, x_12); -lean_dec(x_5); -x_4 = lean_box(0); -x_5 = x_13; -x_6 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg___boxed), 6, 0); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_array_get_size(x_4); +x_4 = lean_alloc_closure((void*)(l_Lean_mkStateFromImportedEntries___rarg___lambda__1), 3, 1); +lean_closure_set(x_4, 0, x_1); +x_5 = lean_array_get_size(x_3); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_nat_dec_lt(x_6, x_5); if (x_7 == 0) { lean_dec(x_5); -lean_dec(x_1); -return x_3; +lean_dec(x_4); +lean_dec(x_3); +return x_2; } else { @@ -8598,7111 +7886,8386 @@ x_8 = lean_nat_dec_le(x_5, x_5); if (x_8 == 0) { lean_dec(x_5); -lean_dec(x_1); -return x_3; +lean_dec(x_4); +lean_dec(x_3); +return x_2; } else { -size_t x_9; size_t x_10; lean_object* x_11; +size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; x_9 = 0; x_10 = lean_usize_of_nat(x_5); lean_dec(x_5); -x_11 = l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg(x_1, x_4, x_9, x_10, x_3); -return x_11; -} -} +x_11 = l_Id_instMonad; +x_12 = l_Array_foldlMUnsafe_fold___rarg(x_11, x_4, x_3, x_9, x_10, x_2); +return x_12; } -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_2, 0); -x_13 = lean_ctor_get(x_2, 1); -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg(x_1, x_12, x_13, lean_box(0), x_14, x_3); -return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg___boxed), 3, 0); -return x_4; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_mkStateFromImportedEntries___rarg), 3, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3478_() { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_2, x_1, x_3); -return x_4; +lean_object* x_1; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__28; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_2, x_1, x_3); -return x_4; +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_apply_1(x_1, x_3); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_2); +lean_ctor_set(x_7, 1, x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_5); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkModuleData___spec__8(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_2) == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) { -return x_1; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_5); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_apply_2(x_8, x_6, x_3); +lean_ctor_set(x_2, 1, x_9); +lean_ctor_set(x_2, 0, x_7); +return x_2; } else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_3 = lean_ctor_get(x_2, 0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_2); lean_inc(x_3); -x_4 = lean_ctor_get(x_2, 1); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_3); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_apply_2(x_13, x_11, x_3); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_12); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 3); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); -x_5 = lean_ctor_get(x_2, 3); -lean_inc(x_5); lean_dec(x_2); -x_6 = l_Lean_RBNode_fold___at_Lean_mkModuleData___spec__8(x_1, x_3); -x_7 = lean_array_push(x_6, x_4); -x_1 = x_7; -x_2 = x_5; -goto _start; -} +x_5 = l_List_reverse___rarg(x_4); +x_6 = lean_apply_1(x_3, x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(lean_object* x_1) { +static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1() { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_3 = l_Lean_RBNode_fold___at_Lean_mkModuleData___spec__8(x_2, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("number of local entries: ", 25, 25); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2() { _start: { -lean_object* x_4; -x_4 = lean_array_push(x_1, x_2); -return x_4; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = lean_array_push(x_1, x_3); -return x_4; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_List_lengthTRAux___rarg(x_2, x_3); +x_5 = l___private_Init_Data_Repr_0__Nat_reprFast(x_4); +x_6 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_6, 0, x_5); +x_7 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } } -static lean_object* _init_l_Lean_mkModuleData___closed__1() { +static lean_object* _init_l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkModuleData___lambda__1___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_mkModuleData___closed__2() { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_mkModuleData___lambda__2___boxed), 3, 0); -return x_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; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_5 = lean_box(0); +x_6 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +lean_inc(x_4); +x_7 = lean_apply_1(x_4, x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_7); +x_9 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_9, 0, x_8); +x_10 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1___boxed), 5, 2); +lean_closure_set(x_10, 0, x_4); +lean_closure_set(x_10, 1, x_5); +lean_inc(x_1); +x_11 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__2), 3, 1); +lean_closure_set(x_11, 0, x_1); +x_12 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__3), 2, 1); +lean_closure_set(x_12, 0, x_1); +x_13 = l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1; +x_14 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_14, 0, x_3); +lean_ctor_set(x_14, 1, x_9); +lean_ctor_set(x_14, 2, x_10); +lean_ctor_set(x_14, 3, x_11); +lean_ctor_set(x_14, 4, x_12); +lean_ctor_set(x_14, 5, x_13); +x_15 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg(x_14, x_2); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_mkModuleData(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; -x_4 = lean_st_ref_get(x_3, x_2); -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_registerSimplePersistentEnvExtension___rarg), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___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; 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; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_6 = lean_ctor_get(x_4, 0); -x_7 = lean_array_size(x_6); -x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(x_1, x_7, x_8, x_6); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = l_Lean_mkModuleData___closed__1; -x_13 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_14 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_12, x_11, x_13); -x_15 = l_Lean_mkModuleData___closed__2; -x_16 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_15, x_11, x_13); -lean_dec(x_11); -x_17 = lean_ctor_get(x_1, 4); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_ctor_get(x_1, 3); -lean_inc(x_19); -lean_dec(x_1); -x_20 = l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(x_19); -x_21 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_21, 0, x_18); -lean_ctor_set(x_21, 1, x_14); -lean_ctor_set(x_21, 2, x_16); -lean_ctor_set(x_21, 3, x_20); -lean_ctor_set(x_21, 4, x_9); -lean_ctor_set(x_4, 0, x_21); -return x_4; -} -else -{ -lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_22 = lean_ctor_get(x_4, 0); -x_23 = lean_ctor_get(x_4, 1); -lean_inc(x_23); -lean_inc(x_22); +lean_object* x_6; +x_6 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); -x_24 = lean_array_size(x_22); -x_25 = 0; -x_26 = l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(x_1, x_24, x_25, x_22); -x_27 = lean_ctor_get(x_1, 1); -lean_inc(x_27); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = l_Lean_mkModuleData___closed__1; -x_30 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_31 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_29, x_28, x_30); -x_32 = l_Lean_mkModuleData___closed__2; -x_33 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_32, x_28, x_30); -lean_dec(x_28); -x_34 = lean_ctor_get(x_1, 4); -lean_inc(x_34); -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -lean_dec(x_34); -x_36 = lean_ctor_get(x_1, 3); -lean_inc(x_36); -lean_dec(x_1); -x_37 = l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(x_36); -x_38 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_38, 0, x_35); -lean_ctor_set(x_38, 1, x_31); -lean_ctor_set(x_38, 2, x_33); -lean_ctor_set(x_38, 3, x_37); -lean_ctor_set(x_38, 4, x_26); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_23); -return x_39; -} +return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___boxed(lean_object* x_1) { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(x_1, x_5, x_6, x_4); +lean_object* x_2; +x_2 = l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4(x_1); lean_dec(x_1); -return x_7; +return x_2; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); +lean_object* x_4; +x_4 = l_Lean_registerSimplePersistentEnvExtension(x_1, x_2, x_3); lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___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_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited___rarg(lean_object* x_1) { _start: { -lean_object* x_7; -x_7 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +x_4 = l_Lean_instInhabitedPersistentEnvExtension(lean_box(0), lean_box(0), lean_box(0), x_3); lean_dec(x_3); -lean_dec(x_2); -return x_7; +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_instInhabited___rarg), 1, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_box(0); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_1); +x_6 = l_Lean_PersistentEnvExtension_getState___rarg(x_5, x_2, x_3); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_getEntries___rarg___boxed), 3, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_mkModuleData___lambda__1(x_1, x_2, x_3); +x_4 = l_Lean_SimplePersistentEnvExtension_getEntries___rarg(x_1, x_2, x_3); lean_dec(x_3); +lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; -x_4 = l_Lean_mkModuleData___lambda__2(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_box(0); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_1); +x_6 = l_Lean_PersistentEnvExtension_getState___rarg(x_5, x_2, x_3); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +return x_7; } } -LEAN_EXPORT lean_object* lean_write_module(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -lean_inc(x_1); -x_4 = l_Lean_mkModuleData(x_1, x_3); -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_dec(x_4); -x_7 = lean_environment_main_module(x_1); -x_8 = lean_save_module_data(x_2, x_7, x_5, x_6); -lean_dec(x_5); -lean_dec(x_7); -lean_dec(x_2); -return x_8; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_getState___rarg___boxed), 3, 0); +return x_3; } } -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; +lean_object* x_4; +x_4 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_name_eq(x_4, x_1); -if (x_6 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -x_2 = x_5; -goto _start; +lean_object* x_4; +x_4 = lean_ctor_get(x_2, 1); +lean_dec(x_4); +lean_ctor_set(x_2, 1, x_1); +return x_2; } else { -uint8_t x_8; -x_8 = 1; -return x_8; +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec(x_2); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_1); +return x_6; +} } } +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_setState___rarg___lambda__1), 2, 1); +lean_closure_set(x_4, 0, x_3); +x_5 = l_Lean_PersistentEnvExtension_modifyState___rarg(x_1, x_2, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_mkExtNameMap___spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState(lean_object* x_1, lean_object* x_2) { _start: { -if (lean_obj_tag(x_2) == 0) +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_setState___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_setState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -return x_1; +lean_object* x_4; +x_4 = l_Lean_SimplePersistentEnvExtension_setState___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___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; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = l_Lean_Name_hash___override(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_apply_1(x_1, x_4); +lean_ctor_set(x_2, 1, x_5); +return x_2; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_7); +lean_inc(x_6); lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = l_Lean_Name_hash___override(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; -} +x_8 = lean_apply_1(x_1, x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_6); +lean_ctor_set(x_9, 1, x_8); +return x_9; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_mkExtNameMap___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) +lean_object* x_4; lean_object* x_5; +x_4 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_modifyState___rarg___lambda__1), 2, 1); +lean_closure_set(x_4, 0, x_3); +x_5 = l_Lean_PersistentEnvExtension_modifyState___rarg(x_1, x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState(lean_object* x_1, lean_object* x_2) { +_start: { -lean_dec(x_2); -lean_dec(x_1); +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_SimplePersistentEnvExtension_modifyState___rarg___boxed), 3, 0); return x_3; } -else +} +LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_modifyState___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_mkExtNameMap___spec__4(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); +lean_object* x_4; +x_4 = l_Lean_SimplePersistentEnvExtension_modifyState___rarg(x_1, x_2, x_3); lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; +return x_4; } } +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3795_() { +_start: +{ +lean_object* x_1; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__28; +return x_1; } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(lean_object* x_1) { +} +static lean_object* _init_l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1() { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_mkExtNameMap___spec__3(x_7, x_1, x_6); -return x_8; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Name_quickLt___boxed), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -if (lean_obj_tag(x_3) == 0) +uint8_t x_7; +x_7 = lean_nat_dec_lt(x_3, x_4); +if (x_7 == 0) { -lean_object* x_4; -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; +lean_dec(x_3); +return x_2; } else { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 1); -x_8 = lean_ctor_get(x_3, 2); -x_9 = lean_name_eq(x_6, x_1); -if (x_9 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1; +lean_inc(x_3); +x_9 = l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(x_1, x_2, x_8, x_3, x_4, lean_box(0), lean_box(0)); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_nat_dec_le(x_4, x_10); +if (x_12 == 0) { -lean_object* x_10; -x_10 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_10); -return x_3; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(x_1, x_11, x_3, x_10, lean_box(0), lean_box(0)); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_10, x_14); +lean_dec(x_10); +x_2 = x_13; +x_3 = x_15; +x_5 = lean_box(0); +x_6 = lean_box(0); +goto _start; } else { -lean_dec(x_7); -lean_dec(x_6); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_1); -return x_3; +lean_dec(x_10); +lean_dec(x_3); +return x_11; } } -else +} +} +LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_3, 0); -x_12 = lean_ctor_get(x_3, 1); -x_13 = lean_ctor_get(x_3, 2); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_3; lean_object* x_4; +x_3 = lean_box(0); +x_4 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_NameSet_empty; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_2 = lean_array_mk(x_1); +x_3 = lean_array_get_size(x_2); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_sub(x_3, x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_3, x_6); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = lean_nat_dec_le(x_6, x_5); +if (x_8 == 0) +{ +lean_object* x_9; +lean_inc(x_5); +x_9 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(x_3, x_2, x_5, x_5, lean_box(0), lean_box(0)); +lean_dec(x_5); lean_dec(x_3); -x_14 = lean_name_eq(x_11, x_1); -if (x_14 == 0) +return x_9; +} +else { -lean_object* x_15; lean_object* x_16; -x_15 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_1, x_2, x_13); -x_16 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_16, 0, x_11); -lean_ctor_set(x_16, 1, x_12); -lean_ctor_set(x_16, 2, x_15); -return x_16; +lean_object* x_10; +x_10 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(x_3, x_2, x_6, x_5, lean_box(0), lean_box(0)); +lean_dec(x_5); +lean_dec(x_3); +return x_10; +} } else { -lean_object* x_17; -lean_dec(x_12); -lean_dec(x_11); -x_17 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_2); -lean_ctor_set(x_17, 2, x_13); -return x_17; +lean_dec(x_5); +lean_dec(x_3); +return x_2; } } } +static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__1), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___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) { +static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__2() { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_4, 1); -x_11 = lean_nat_dec_lt(x_6, x_10); -if (x_11 == 0) -{ -lean_object* x_12; -lean_dec(x_6); -lean_dec(x_2); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_9); -return x_12; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__2___boxed), 1, 0); +return x_1; } -else -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_array_fget(x_1, x_6); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = !lean_is_exclusive(x_5); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; lean_object* x_31; uint8_t x_32; -x_16 = lean_ctor_get(x_5, 0); -x_17 = lean_ctor_get(x_5, 1); -x_18 = lean_array_get_size(x_17); -x_19 = l_Lean_Name_hash___override(x_14); -x_20 = 32; -x_21 = lean_uint64_shift_right(x_19, x_20); -x_22 = lean_uint64_xor(x_19, x_21); -x_23 = 16; -x_24 = lean_uint64_shift_right(x_22, x_23); -x_25 = lean_uint64_xor(x_22, x_24); -x_26 = lean_uint64_to_usize(x_25); -x_27 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_28 = 1; -x_29 = lean_usize_sub(x_27, x_28); -x_30 = lean_usize_land(x_26, x_29); -x_31 = lean_array_uget(x_17, x_30); -x_32 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_14, x_31); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_33 = lean_unsigned_to_nat(1u); -x_34 = lean_nat_add(x_16, x_33); -lean_dec(x_16); -lean_inc(x_6); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_14); -lean_ctor_set(x_35, 1, x_6); -lean_ctor_set(x_35, 2, x_31); -x_36 = lean_array_uset(x_17, x_30, x_35); -x_37 = lean_unsigned_to_nat(4u); -x_38 = lean_nat_mul(x_34, x_37); -x_39 = lean_unsigned_to_nat(3u); -x_40 = lean_nat_div(x_38, x_39); -lean_dec(x_38); -x_41 = lean_array_get_size(x_36); -x_42 = lean_nat_dec_le(x_40, x_41); -lean_dec(x_41); -lean_dec(x_40); -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(x_36); -lean_ctor_set(x_5, 1, x_43); -lean_ctor_set(x_5, 0, x_34); -x_44 = lean_ctor_get(x_4, 2); -x_45 = lean_nat_add(x_6, x_44); -lean_dec(x_6); -x_6 = x_45; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; } -else +static lean_object* _init_l_Lean_mkTagDeclarationExtension___closed__3() { +_start: { -lean_object* x_47; lean_object* x_48; -lean_ctor_set(x_5, 1, x_36); -lean_ctor_set(x_5, 0, x_34); -x_47 = lean_ctor_get(x_4, 2); -x_48 = lean_nat_add(x_6, x_47); -lean_dec(x_6); -x_6 = x_48; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkTagDeclarationExtension___lambda__3), 1, 0); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_inc(x_2); -x_50 = lean_array_uset(x_17, x_30, x_2); -lean_inc(x_6); -x_51 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_14, x_6, x_31); -x_52 = lean_array_uset(x_50, x_30, x_51); -lean_ctor_set(x_5, 1, x_52); -x_53 = lean_ctor_get(x_4, 2); -x_54 = lean_nat_add(x_6, x_53); -lean_dec(x_6); -x_6 = x_54; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Lean_mkTagDeclarationExtension___closed__1; +x_4 = l_Lean_mkTagDeclarationExtension___closed__2; +x_5 = l_Lean_mkTagDeclarationExtension___closed__3; +x_6 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set(x_6, 3, x_5); +x_7 = l_Lean_registerSimplePersistentEnvExtension___rarg(x_6, x_2); +return x_7; } } -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint64_t x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; uint64_t x_65; size_t x_66; size_t x_67; size_t x_68; size_t x_69; size_t x_70; lean_object* x_71; uint8_t x_72; -x_56 = lean_ctor_get(x_5, 0); -x_57 = lean_ctor_get(x_5, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_5); -x_58 = lean_array_get_size(x_57); -x_59 = l_Lean_Name_hash___override(x_14); -x_60 = 32; -x_61 = lean_uint64_shift_right(x_59, x_60); -x_62 = lean_uint64_xor(x_59, x_61); -x_63 = 16; -x_64 = lean_uint64_shift_right(x_62, x_63); -x_65 = lean_uint64_xor(x_62, x_64); -x_66 = lean_uint64_to_usize(x_65); -x_67 = lean_usize_of_nat(x_58); -lean_dec(x_58); -x_68 = 1; -x_69 = lean_usize_sub(x_67, x_68); -x_70 = lean_usize_land(x_66, x_69); -x_71 = lean_array_uget(x_57, x_70); -x_72 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_14, x_71); -if (x_72 == 0) +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; -x_73 = lean_unsigned_to_nat(1u); -x_74 = lean_nat_add(x_56, x_73); -lean_dec(x_56); -lean_inc(x_6); -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_14); -lean_ctor_set(x_75, 1, x_6); -lean_ctor_set(x_75, 2, x_71); -x_76 = lean_array_uset(x_57, x_70, x_75); -x_77 = lean_unsigned_to_nat(4u); -x_78 = lean_nat_mul(x_74, x_77); -x_79 = lean_unsigned_to_nat(3u); -x_80 = lean_nat_div(x_78, x_79); -lean_dec(x_78); -x_81 = lean_array_get_size(x_76); -x_82 = lean_nat_dec_le(x_80, x_81); -lean_dec(x_81); -lean_dec(x_80); -if (x_82 == 0) +lean_object* x_7; +x_7 = l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension___lambda__2___boxed(lean_object* x_1) { +_start: { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_83 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(x_76); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_74); -lean_ctor_set(x_84, 1, x_83); -x_85 = lean_ctor_get(x_4, 2); -x_86 = lean_nat_add(x_6, x_85); -lean_dec(x_6); -x_5 = x_84; -x_6 = x_86; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_object* x_2; +x_2 = l_Lean_mkTagDeclarationExtension___lambda__2(x_1); +lean_dec(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_TagDeclarationExtension_instInhabited___closed__1() { +_start: { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_74); -lean_ctor_set(x_88, 1, x_76); -x_89 = lean_ctor_get(x_4, 2); -x_90 = lean_nat_add(x_6, x_89); -lean_dec(x_6); -x_5 = x_88; -x_6 = x_90; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_NameSet_instInhabited; +x_2 = l_Lean_SimplePersistentEnvExtension_instInhabited___rarg(x_1); +return x_2; } } -else +static lean_object* _init_l_Lean_TagDeclarationExtension_instInhabited() { +_start: { -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_inc(x_2); -x_92 = lean_array_uset(x_57, x_70, x_2); -lean_inc(x_6); -x_93 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_14, x_6, x_71); -x_94 = lean_array_uset(x_92, x_70, x_93); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_56); -lean_ctor_set(x_95, 1, x_94); -x_96 = lean_ctor_get(x_4, 2); -x_97 = lean_nat_add(x_6, x_96); -lean_dec(x_6); -x_5 = x_95; -x_6 = x_97; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_object* x_1; +x_1 = l_Lean_TagDeclarationExtension_instInhabited___closed__1; +return x_1; +} } +LEAN_EXPORT lean_object* l_panic___at_Lean_TagDeclarationExtension_tag___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_panic_fn(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkExtNameMap(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__2() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_3 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; -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_dec(x_4); -x_7 = lean_box(0); -x_8 = lean_array_get_size(x_5); -x_9 = lean_unsigned_to_nat(1u); -lean_inc(x_1); -x_10 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_8); -lean_ctor_set(x_10, 2, x_9); -x_11 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; -x_12 = l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spec__6(x_5, x_7, x_10, x_10, x_11, x_1, lean_box(0), lean_box(0), x_6); -lean_dec(x_10); -lean_dec(x_5); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("env.getModuleIdxFor\? declName |>.isNone -- See comment at `TagDeclarationExtension`\n ", 86, 86); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__3() { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_TagDeclarationExtension_tag___closed__1; +x_2 = l_Lean_TagDeclarationExtension_tag___closed__2; +x_3 = lean_string_append(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___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) { +static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__4() { _start: { -lean_object* x_10; -x_10 = l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.TagDeclarationExtension.tag", 32, 32); +return x_1; } } -LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_TagDeclarationExtension_tag___closed__5() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_3 = lean_array_get_size(x_1); -x_4 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_5 = lean_mk_array(x_3, x_4); -x_6 = !lean_is_exclusive(x_2); -if (x_6 == 0) -{ -lean_object* x_7; -x_7 = lean_ctor_get(x_2, 0); -lean_dec(x_7); -lean_ctor_set(x_2, 0, x_5); -return x_2; -} -else -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_2, 1); -lean_inc(x_8); -lean_dec(x_2); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_5); -lean_ctor_set(x_9, 1, x_8); -return x_9; -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; +x_2 = l_Lean_TagDeclarationExtension_tag___closed__4; +x_3 = lean_unsigned_to_nat(729u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l_Lean_TagDeclarationExtension_tag___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_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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_EXPORT lean_object* l_Lean_TagDeclarationExtension_tag(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_4, x_3); -if (x_7 == 0) +lean_object* x_4; +x_4 = l_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); +if (lean_obj_tag(x_4) == 0) { -lean_object* x_8; -lean_dec(x_1); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_5); -lean_ctor_set(x_8, 1, x_6); -return x_8; +lean_object* x_5; +x_5 = l_Lean_PersistentEnvExtension_addEntry___rarg(x_1, x_2, x_3); +return x_5; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; -x_9 = lean_ctor_get(x_2, 0); -x_10 = lean_array_uget(x_9, x_4); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -lean_inc(x_1); -x_12 = lean_alloc_closure((void*)(l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1___boxed), 2, 1); -lean_closure_set(x_12, 0, x_1); -x_13 = l_Lean_EnvExtension_modifyState___rarg(x_11, x_5, x_12); -lean_dec(x_11); -x_14 = 1; -x_15 = lean_usize_add(x_4, x_14); -x_4 = x_15; -x_5 = x_13; -goto _start; +lean_object* x_6; lean_object* x_7; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_6 = l_Lean_TagDeclarationExtension_tag___closed__5; +x_7 = l_panic___at_Lean_TagDeclarationExtension_tag___spec__1(x_2, x_6); +return x_7; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -if (lean_obj_tag(x_2) == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_nat_add(x_3, x_4); +x_7 = lean_unsigned_to_nat(2u); +x_8 = lean_nat_div(x_6, x_7); +lean_dec(x_6); +x_9 = lean_array_fget(x_1, x_8); +x_10 = l_Lean_Name_quickLt(x_9, x_2); +if (x_10 == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +uint8_t x_11; +lean_dec(x_4); +x_11 = l_Lean_Name_quickLt(x_2, x_9); +lean_dec(x_9); +if (x_11 == 0) +{ +uint8_t x_12; +lean_dec(x_8); +lean_dec(x_3); +x_12 = 1; +return x_12; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_name_eq(x_4, x_1); -if (x_7 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_8, x_13); +if (x_14 == 0) { -x_2 = x_6; +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_8, x_15); +lean_dec(x_8); +x_17 = lean_nat_dec_lt(x_16, x_3); +if (x_17 == 0) +{ +x_4 = x_16; +x_5 = lean_box(0); goto _start; } else { -lean_object* x_9; -lean_inc(x_5); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_5); -return x_9; +uint8_t x_19; +lean_dec(x_16); +lean_dec(x_3); +x_19 = 0; +return x_19; +} } +else +{ +uint8_t x_20; +lean_dec(x_8); +lean_dec(x_3); +x_20 = 0; +return x_20; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_dec(x_9); +lean_dec(x_3); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_8, x_21); +lean_dec(x_8); +x_23 = lean_nat_dec_le(x_22, x_4); +if (x_23 == 0) { -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_array_set(x_5, x_1, x_2); -lean_ctor_set(x_3, 0, x_6); -return x_3; +uint8_t x_24; +lean_dec(x_22); +lean_dec(x_4); +x_24 = 0; +return x_24; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_ctor_get(x_3, 0); -x_8 = lean_ctor_get(x_3, 1); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_3); -x_9 = lean_array_set(x_7, x_1, x_2); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; +x_3 = x_22; +x_5 = lean_box(0); +goto _start; } } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___closed__1() { +} +static lean_object* _init_l_Lean_TagDeclarationExtension_isTagged___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_instInhabitedEnvExtensionState; -x_2 = l_Lean_instInhabitedPersistentEnvExtension(lean_box(0), lean_box(0), lean_box(0), x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_NameSet_instInhabited; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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_EXPORT uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_lt(x_8, x_7); -if (x_11 == 0) -{ -lean_object* x_12; -lean_dec(x_3); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_9); -lean_ctor_set(x_12, 1, x_10); -return x_12; +lean_object* x_4; +x_4 = l_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = l_Lean_NameSet_instInhabited; +x_6 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_5, x_1, x_2); +x_7 = l_Lean_NameSet_contains(x_6, x_3); +lean_dec(x_6); +return x_7; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; -x_13 = lean_array_uget(x_6, x_8); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_2, 1); -x_17 = lean_array_get_size(x_16); -x_18 = l_Lean_Name_hash___override(x_14); -x_19 = 32; -x_20 = lean_uint64_shift_right(x_18, x_19); -x_21 = lean_uint64_xor(x_18, x_20); -x_22 = 16; -x_23 = lean_uint64_shift_right(x_21, x_22); -x_24 = lean_uint64_xor(x_21, x_23); -x_25 = lean_uint64_to_usize(x_24); -x_26 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_27 = 1; -x_28 = lean_usize_sub(x_26, x_27); -x_29 = lean_usize_land(x_25, x_28); -x_30 = lean_array_uget(x_16, x_29); -x_31 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(x_14, x_30); -lean_dec(x_30); -lean_dec(x_14); -if (lean_obj_tag(x_31) == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_8 = lean_ctor_get(x_4, 0); +lean_inc(x_8); +lean_dec(x_4); +x_9 = l_Lean_TagDeclarationExtension_isTagged___closed__1; +x_10 = l_Lean_PersistentEnvExtension_getModuleEntries___rarg(x_9, x_1, x_2, x_8); +lean_dec(x_8); +x_11 = lean_array_get_size(x_10); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_sub(x_11, x_12); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_lt(x_14, x_11); +lean_dec(x_11); +if (x_15 == 0) { -size_t x_32; -lean_dec(x_15); -x_32 = lean_usize_add(x_8, x_27); -x_8 = x_32; -goto _start; +uint8_t x_16; +lean_dec(x_13); +lean_dec(x_10); +x_16 = 0; +return x_16; } else { -lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_31, 0); -lean_inc(x_34); -lean_dec(x_31); -x_35 = lean_array_get_size(x_1); -x_36 = lean_nat_dec_lt(x_34, x_35); -lean_dec(x_35); -lean_inc(x_3); -x_37 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1___boxed), 3, 2); -lean_closure_set(x_37, 0, x_3); -lean_closure_set(x_37, 1, x_15); -if (x_36 == 0) +uint8_t x_17; +x_17 = lean_nat_dec_le(x_14, x_13); +if (x_17 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; -lean_dec(x_34); -x_38 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___closed__1; -x_39 = l_outOfBounds___rarg(x_38); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -lean_dec(x_39); -x_41 = l_Lean_EnvExtension_modifyState___rarg(x_40, x_9, x_37); -lean_dec(x_40); -x_42 = lean_usize_add(x_8, x_27); -x_8 = x_42; -x_9 = x_41; -goto _start; +uint8_t x_18; +lean_dec(x_13); +lean_dec(x_10); +x_18 = 0; +return x_18; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; -x_44 = lean_array_fget(x_1, x_34); -lean_dec(x_34); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -lean_dec(x_44); -x_46 = l_Lean_EnvExtension_modifyState___rarg(x_45, x_9, x_37); -lean_dec(x_45); -x_47 = lean_usize_add(x_8, x_27); -x_8 = x_47; -x_9 = x_46; -goto _start; +uint8_t x_19; +x_19 = l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1(x_10, x_3, x_14, x_13, lean_box(0)); +lean_dec(x_10); +return x_19; } } } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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_EXPORT lean_object* l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_5, 1); -x_12 = lean_nat_dec_lt(x_7, x_11); -if (x_12 == 0) -{ -lean_object* x_13; -lean_dec(x_7); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_6); -lean_ctor_set(x_13, 1, x_10); -return x_13; +uint8_t x_6; lean_object* x_7; +x_6 = l_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; } -else +} +LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_isTagged___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_14 = lean_array_fget(x_1, x_7); -x_15 = lean_ctor_get(x_14, 4); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_box(0); -x_17 = lean_array_size(x_15); -x_18 = 0; -lean_inc(x_7); -x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3(x_2, x_3, x_7, x_15, x_16, x_15, x_17, x_18, x_6, x_10); -lean_dec(x_15); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_5, 2); -x_23 = lean_nat_add(x_7, x_22); -lean_dec(x_7); -x_6 = x_20; -x_7 = x_23; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_10 = x_21; -goto _start; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_TagDeclarationExtension_isTagged(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; } } +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3958_() { +_start: +{ +lean_object* x_1; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__28; +return x_1; +} } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_setImportedEntries(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; lean_object* x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_5 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; -x_6 = lean_st_ref_get(x_5, x_4); -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_array_get_size(x_7); -lean_inc(x_3); -lean_inc(x_7); -x_10 = l_Array_toSubarray___rarg(x_7, x_3, x_9); -x_11 = lean_ctor_get(x_10, 2); -lean_inc(x_11); -x_12 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -x_14 = lean_usize_of_nat(x_13); -lean_dec(x_13); -lean_inc(x_2); -x_15 = l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1(x_2, x_10, x_12, x_14, x_1, x_8); -lean_dec(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 = l_Lean_mkExtNameMap(x_3, x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_array_get_size(x_2); -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_23); -x_25 = l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4(x_2, x_7, x_19, x_24, x_24, x_16, x_22, lean_box(0), lean_box(0), x_20); -lean_dec(x_24); -lean_dec(x_19); -lean_dec(x_7); -lean_dec(x_2); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +if (lean_obj_tag(x_2) == 0) { -return x_25; +return x_1; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_25, 0); -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_25); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_3; 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 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ctor_get(x_2, 3); +x_7 = l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(x_1, x_3); +lean_inc(x_5); +lean_inc(x_4); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_4); +lean_ctor_set(x_8, 1, x_5); +x_9 = lean_array_push(x_7, x_8); +x_1 = x_9; +x_2 = x_6; +goto _start; } } } -LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1(x_1, x_2); -lean_dec(x_1); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_3 = l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1(lean_object* x_1) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_9 = l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1(x_1, x_2, x_7, x_8, x_5, x_6); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); lean_dec(x_2); -return x_9; +x_5 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_1, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_mkMapDeclarationExtension___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_mkMapDeclarationExtension___rarg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkMapDeclarationExtension___rarg___lambda__2), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_box(0); +x_4 = lean_alloc_closure((void*)(l_Lean_mkMapDeclarationExtension___rarg___lambda__1___boxed), 4, 1); +lean_closure_set(x_4, 0, x_3); +x_5 = l_Lean_mkMapDeclarationExtension___rarg___closed__1; +x_6 = l_Lean_mkMapDeclarationExtension___rarg___closed__2; +x_7 = lean_alloc_closure((void*)(l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed), 1, 0); +x_8 = l_Lean_instInhabitedPersistentEnvExtension___closed__4; +x_9 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_5); +lean_ctor_set(x_9, 2, x_4); +lean_ctor_set(x_9, 3, x_6); +lean_ctor_set(x_9, 4, x_7); +lean_ctor_set(x_9, 5, x_8); +x_10 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg(x_9, x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_mkMapDeclarationExtension___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(x_1, x_2); +x_3 = l_Lean_RBNode_fold___at_Lean_mkMapDeclarationExtension___spec__2___rarg(x_1, x_2); lean_dec(x_2); -lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg___boxed(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1(x_1, x_2, x_3); +lean_object* x_2; +x_2 = l_Lean_RBMap_toArray___at_Lean_mkMapDeclarationExtension___spec__1___rarg(x_1); lean_dec(x_1); -return x_4; +return x_2; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_12 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_12, x_9, x_10); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_5; +x_5 = l_Lean_mkMapDeclarationExtension___rarg___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_13; +return x_5; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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) { +static lean_object* _init_l_Lean_MapDeclarationExtension_instInhabited___closed__1() { _start: { -lean_object* x_11; -x_11 = l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = l_Lean_instInhabitedPersistentEnvExtension(lean_box(0), lean_box(0), lean_box(0), x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_instInhabited(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_MapDeclarationExtension_instInhabited___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_update_env_attributes(x_1, x_2); +x_3 = lean_panic_fn(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_getNumBuiltinAttributes___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__1() { _start: { -lean_object* x_2; -x_2 = lean_get_num_attributes(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("env.getModuleIdxFor\? declName |>.isNone -- See comment at `MapDeclarationExtension`\n ", 86, 86); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__2() { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_unsigned_to_nat(1u); -x_8 = lean_nat_add(x_1, x_7); -x_9 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(x_2, x_3, x_8, x_4, x_6); -lean_dec(x_8); -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_TagDeclarationExtension_tag___closed__1; +x_2 = l_Lean_MapDeclarationExtension_insert___rarg___closed__1; +x_3 = lean_string_append(x_1, x_2); +return x_3; } } -static lean_object* _init_l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1() { +static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_instInhabitedEnvExtensionState; -x_2 = l_Lean_instInhabitedPersistentEnvExtensionState___rarg(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.MapDeclarationExtension.insert", 35, 35); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_MapDeclarationExtension_insert___rarg___closed__4() { _start: { -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; -x_7 = lean_st_ref_get(x_6, x_5); -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__1; +x_2 = l_Lean_MapDeclarationExtension_insert___rarg___closed__3; +x_3 = lean_unsigned_to_nat(760u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l_Lean_MapDeclarationExtension_insert___rarg___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_ctor_get(x_7, 1); -x_11 = lean_array_get_size(x_9); -x_12 = lean_nat_dec_lt(x_3, x_11); -lean_dec(x_11); -if (x_12 == 0) +lean_object* x_5; +x_5 = l_Lean_Environment_getModuleIdxFor_x3f(x_2, x_3); +if (lean_obj_tag(x_5) == 0) { -lean_dec(x_9); -lean_dec(x_2); -lean_dec(x_1); -lean_ctor_set(x_7, 0, x_4); +lean_object* x_6; lean_object* x_7; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_4); +x_7 = l_Lean_PersistentEnvExtension_addEntry___rarg(x_1, x_2, x_6); return x_7; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_free_object(x_7); -x_13 = lean_array_fget(x_9, x_3); -lean_dec(x_9); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1; -x_16 = l_Lean_EnvExtension_getState___rarg(x_15, x_14, x_4); -x_17 = lean_st_ref_get(x_6, x_10); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +lean_object* x_8; lean_object* x_9; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_8 = l_Lean_MapDeclarationExtension_insert___rarg___closed__4; +x_9 = l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(x_2, x_8); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_insert(lean_object* x_1) { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -x_21 = lean_array_get_size(x_19); -lean_dec(x_19); -x_22 = lean_get_num_attributes(x_20); -if (lean_obj_tag(x_22) == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_insert___rarg), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_13, 2); -lean_inc(x_25); -lean_dec(x_13); -x_26 = !lean_is_exclusive(x_16); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_16, 0); -x_28 = lean_ctor_get(x_16, 1); -lean_dec(x_28); -lean_inc(x_2); -lean_inc(x_4); -lean_ctor_set(x_17, 1, x_2); -lean_ctor_set(x_17, 0, x_4); -lean_inc(x_27); -x_29 = lean_apply_3(x_25, x_27, x_17, x_24); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -lean_ctor_set(x_16, 1, x_30); -x_32 = l_Lean_EnvExtension_setState___rarg(x_14, x_4, x_16); -lean_dec(x_14); -x_33 = l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(x_32, x_31); -if (lean_obj_tag(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 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_st_ref_get(x_6, x_35); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_get_num_attributes(x_38); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_array_get_size(x_37); -lean_dec(x_37); -x_43 = lean_nat_dec_lt(x_21, x_42); -lean_dec(x_42); -if (x_43 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_nat_add(x_3, x_4); +x_7 = lean_unsigned_to_nat(2u); +x_8 = lean_nat_div(x_6, x_7); +lean_dec(x_6); +x_9 = lean_array_fget(x_1, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 0); +x_12 = l_Lean_Name_quickLt(x_10, x_11); +if (x_12 == 0) { -uint8_t x_44; -x_44 = lean_nat_dec_lt(x_23, x_40); -lean_dec(x_40); -lean_dec(x_23); -if (x_44 == 0) +uint8_t x_13; +lean_dec(x_4); +x_13 = l_Lean_Name_quickLt(x_11, x_10); +lean_dec(x_10); +if (x_13 == 0) { -lean_object* x_45; lean_object* x_46; -lean_dec(x_21); -x_45 = lean_box(0); -x_46 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_34, x_45, x_41); -return x_46; +lean_object* x_14; +lean_dec(x_8); +lean_dec(x_3); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_9); +return x_14; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -lean_inc(x_1); -x_47 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_34, x_1, x_21, x_41); -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -lean_dec(x_47); -x_50 = lean_update_env_attributes(x_48, x_49); -if (lean_obj_tag(x_50) == 0) +lean_object* x_15; uint8_t x_16; +lean_dec(x_9); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_eq(x_8, x_15); +if (x_16 == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); -x_53 = lean_box(0); -x_54 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_51, x_53, x_52); -return x_54; +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_8, x_17); +lean_dec(x_8); +x_19 = lean_nat_dec_lt(x_18, x_3); +if (x_19 == 0) +{ +x_4 = x_18; +x_5 = lean_box(0); +goto _start; } else { -uint8_t x_55; -lean_dec(x_2); -lean_dec(x_1); -x_55 = !lean_is_exclusive(x_50); -if (x_55 == 0) -{ -return x_50; +lean_object* x_21; +lean_dec(x_18); +lean_dec(x_3); +x_21 = lean_box(0); +return x_21; +} } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_50, 0); -x_57 = lean_ctor_get(x_50, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_50); -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_22; +lean_dec(x_8); +lean_dec(x_3); +x_22 = lean_box(0); +return x_22; } } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_40); -lean_dec(x_23); -lean_inc(x_1); -x_59 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_34, x_1, x_21, x_41); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_update_env_attributes(x_60, x_61); -if (lean_obj_tag(x_62) == 0) +lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_8, x_23); +lean_dec(x_8); +x_25 = lean_nat_dec_le(x_24, x_4); +if (x_25 == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_box(0); -x_66 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_63, x_65, x_64); -return x_66; +lean_object* x_26; +lean_dec(x_24); +lean_dec(x_4); +x_26 = lean_box(0); +return x_26; } else { -uint8_t x_67; -lean_dec(x_2); -lean_dec(x_1); -x_67 = !lean_is_exclusive(x_62); -if (x_67 == 0) -{ -return x_62; +x_3 = x_24; +x_5 = lean_box(0); +goto _start; } -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_62, 0); -x_69 = lean_ctor_get(x_62, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_62); -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_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg___boxed), 5, 0); +return x_2; +} } -else +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_71; -lean_dec(x_37); -lean_dec(x_34); -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_2); -lean_dec(x_1); -x_71 = !lean_is_exclusive(x_39); -if (x_71 == 0) +lean_object* x_5; +x_5 = l_Lean_Environment_getModuleIdxFor_x3f(x_3, x_4); +if (lean_obj_tag(x_5) == 0) { -return x_39; +lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_1); +x_6 = lean_box(0); +x_7 = l_Lean_PersistentEnvExtension_getState___rarg(x_6, x_2, x_3); +x_8 = l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(x_7, x_4); +lean_dec(x_4); +lean_dec(x_7); +return x_8; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_39, 0); -x_73 = lean_ctor_get(x_39, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_39); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; -} -} +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_9 = lean_ctor_get(x_5, 0); +lean_inc(x_9); +lean_dec(x_5); +x_10 = lean_box(0); +x_11 = l_Lean_PersistentEnvExtension_getModuleEntries___rarg(x_10, x_2, x_3, x_9); +lean_dec(x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_4); +lean_ctor_set(x_12, 1, x_1); +x_13 = lean_array_get_size(x_11); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_13, x_14); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_lt(x_16, x_13); +lean_dec(x_13); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +x_18 = lean_box(0); +return x_18; } else { -uint8_t x_75; -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_2); -lean_dec(x_1); -x_75 = !lean_is_exclusive(x_33); -if (x_75 == 0) +uint8_t x_19; +x_19 = lean_nat_dec_le(x_16, x_15); +if (x_19 == 0) { -return x_33; +lean_object* x_20; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +x_20 = lean_box(0); +return x_20; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_33, 0); -x_77 = lean_ctor_get(x_33, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_33); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -return x_78; -} -} +lean_object* x_21; +x_21 = l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg(x_11, x_12, x_16, x_15, lean_box(0)); +lean_dec(x_12); +lean_dec(x_11); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; +x_22 = lean_box(0); +return x_22; } else { -uint8_t x_79; -lean_free_object(x_16); -lean_dec(x_27); -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_14); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_79 = !lean_is_exclusive(x_29); -if (x_79 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) { -return x_29; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 0); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +lean_ctor_set(x_21, 0, x_25); +return x_21; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_29, 0); -x_81 = lean_ctor_get(x_29, 1); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_29); -x_82 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -return x_82; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_21, 0); +lean_inc(x_26); +lean_dec(x_21); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +return x_28; } } } -else +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f(lean_object* x_1) { +_start: { -lean_object* x_83; lean_object* x_84; -x_83 = lean_ctor_get(x_16, 0); -lean_inc(x_83); -lean_dec(x_16); -lean_inc(x_2); -lean_inc(x_4); -lean_ctor_set(x_17, 1, x_2); -lean_ctor_set(x_17, 0, x_4); -lean_inc(x_83); -x_84 = lean_apply_3(x_25, x_83, x_17, x_24); -if (lean_obj_tag(x_84) == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_find_x3f___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_84, 1); -lean_inc(x_86); -lean_dec(x_84); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_83); -lean_ctor_set(x_87, 1, x_85); -x_88 = l_Lean_EnvExtension_setState___rarg(x_14, x_4, x_87); -lean_dec(x_14); -x_89 = l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(x_88, x_86); -if (lean_obj_tag(x_89) == 0) +lean_object* x_6; +x_6 = l_Array_binSearchAux___at_Lean_MapDeclarationExtension_find_x3f___spec__1___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -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_st_ref_get(x_6, x_91); -x_93 = lean_ctor_get(x_92, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_92, 1); -lean_inc(x_94); -lean_dec(x_92); -x_95 = lean_get_num_attributes(x_94); -if (lean_obj_tag(x_95) == 0) +lean_object* x_5; +x_5 = l_Lean_MapDeclarationExtension_find_x3f___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -lean_dec(x_95); -x_98 = lean_array_get_size(x_93); -lean_dec(x_93); -x_99 = lean_nat_dec_lt(x_21, x_98); -lean_dec(x_98); -if (x_99 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_nat_add(x_3, x_4); +x_7 = lean_unsigned_to_nat(2u); +x_8 = lean_nat_div(x_6, x_7); +lean_dec(x_6); +x_9 = lean_array_fget(x_1, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_ctor_get(x_2, 0); +x_12 = l_Lean_Name_quickLt(x_10, x_11); +if (x_12 == 0) { -uint8_t x_100; -x_100 = lean_nat_dec_lt(x_23, x_96); -lean_dec(x_96); -lean_dec(x_23); -if (x_100 == 0) +uint8_t x_13; +lean_dec(x_4); +x_13 = l_Lean_Name_quickLt(x_11, x_10); +lean_dec(x_10); +if (x_13 == 0) { -lean_object* x_101; lean_object* x_102; -lean_dec(x_21); -x_101 = lean_box(0); -x_102 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_90, x_101, x_97); -return x_102; +uint8_t x_14; +lean_dec(x_8); +lean_dec(x_3); +x_14 = 1; +return x_14; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_inc(x_1); -x_103 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_90, x_1, x_21, x_97); -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_103, 1); -lean_inc(x_105); -lean_dec(x_103); -x_106 = lean_update_env_attributes(x_104, x_105); -if (lean_obj_tag(x_106) == 0) +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_eq(x_8, x_15); +if (x_16 == 0) { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); -lean_dec(x_106); -x_109 = lean_box(0); -x_110 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_107, x_109, x_108); -return x_110; +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_8, x_17); +lean_dec(x_8); +x_19 = lean_nat_dec_lt(x_18, x_3); +if (x_19 == 0) +{ +x_4 = x_18; +x_5 = lean_box(0); +goto _start; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -lean_dec(x_2); -lean_dec(x_1); -x_111 = lean_ctor_get(x_106, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_106, 1); -lean_inc(x_112); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_113 = x_106; -} else { - lean_dec_ref(x_106); - x_113 = lean_box(0); +uint8_t x_21; +lean_dec(x_18); +lean_dec(x_3); +x_21 = 0; +return x_21; } -if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(1, 2, 0); -} else { - x_114 = x_113; } -lean_ctor_set(x_114, 0, x_111); -lean_ctor_set(x_114, 1, x_112); -return x_114; +else +{ +uint8_t x_22; +lean_dec(x_8); +lean_dec(x_3); +x_22 = 0; +return x_22; } } } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -lean_dec(x_96); -lean_dec(x_23); -lean_inc(x_1); -x_115 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_90, x_1, x_21, x_97); -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_115, 1); -lean_inc(x_117); -lean_dec(x_115); -x_118 = lean_update_env_attributes(x_116, x_117); -if (lean_obj_tag(x_118) == 0) +lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_dec(x_10); +lean_dec(x_3); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_8, x_23); +lean_dec(x_8); +x_25 = lean_nat_dec_le(x_24, x_4); +if (x_25 == 0) { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_119 = lean_ctor_get(x_118, 0); -lean_inc(x_119); -x_120 = lean_ctor_get(x_118, 1); -lean_inc(x_120); -lean_dec(x_118); -x_121 = lean_box(0); -x_122 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_119, x_121, x_120); -return x_122; +uint8_t x_26; +lean_dec(x_24); +lean_dec(x_4); +x_26 = 0; +return x_26; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -lean_dec(x_2); -lean_dec(x_1); -x_123 = lean_ctor_get(x_118, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_118, 1); -lean_inc(x_124); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - x_125 = x_118; -} else { - lean_dec_ref(x_118); - x_125 = lean_box(0); -} -if (lean_is_scalar(x_125)) { - x_126 = lean_alloc_ctor(1, 2, 0); -} else { - x_126 = x_125; +x_3 = x_24; +x_5 = lean_box(0); +goto _start; } -lean_ctor_set(x_126, 0, x_123); -lean_ctor_set(x_126, 1, x_124); -return x_126; } } } -else +LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1(lean_object* x_1) { +_start: { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; -lean_dec(x_93); -lean_dec(x_90); -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_2); -lean_dec(x_1); -x_127 = lean_ctor_get(x_95, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_95, 1); -lean_inc(x_128); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_129 = x_95; -} else { - lean_dec_ref(x_95); - x_129 = lean_box(0); -} -if (lean_is_scalar(x_129)) { - x_130 = lean_alloc_ctor(1, 2, 0); -} else { - x_130 = x_129; -} -lean_ctor_set(x_130, 0, x_127); -lean_ctor_set(x_130, 1, x_128); -return x_130; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg___boxed), 5, 0); +return x_2; } } -else +LEAN_EXPORT uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_2); +lean_object* x_5; +x_5 = l_Lean_Environment_getModuleIdxFor_x3f(x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_dec(x_1); -x_131 = lean_ctor_get(x_89, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_89, 1); -lean_inc(x_132); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_133 = x_89; -} else { - lean_dec_ref(x_89); - x_133 = lean_box(0); -} -if (lean_is_scalar(x_133)) { - x_134 = lean_alloc_ctor(1, 2, 0); -} else { - x_134 = x_133; +x_6 = lean_box(0); +x_7 = l_Lean_PersistentEnvExtension_getState___rarg(x_6, x_2, x_3); +x_8 = l_Lean_NameMap_contains___rarg(x_7, x_4); +lean_dec(x_4); +lean_dec(x_7); +return x_8; } -lean_ctor_set(x_134, 0, x_131); -lean_ctor_set(x_134, 1, x_132); -return x_134; +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_9 = lean_ctor_get(x_5, 0); +lean_inc(x_9); +lean_dec(x_5); +x_10 = lean_box(0); +x_11 = l_Lean_PersistentEnvExtension_getModuleEntries___rarg(x_10, x_2, x_3, x_9); +lean_dec(x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_4); +lean_ctor_set(x_12, 1, x_1); +x_13 = lean_array_get_size(x_11); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_13, x_14); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_lt(x_16, x_13); +lean_dec(x_13); +if (x_17 == 0) +{ +uint8_t x_18; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +x_18 = 0; +return x_18; } +else +{ +uint8_t x_19; +x_19 = lean_nat_dec_le(x_16, x_15); +if (x_19 == 0) +{ +uint8_t x_20; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +x_20 = 0; +return x_20; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -lean_dec(x_83); -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_14); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_135 = lean_ctor_get(x_84, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_84, 1); -lean_inc(x_136); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_137 = x_84; -} else { - lean_dec_ref(x_84); - x_137 = lean_box(0); +uint8_t x_21; +x_21 = l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(x_11, x_12, x_16, x_15, lean_box(0)); +lean_dec(x_12); +lean_dec(x_11); +return x_21; } -if (lean_is_scalar(x_137)) { - x_138 = lean_alloc_ctor(1, 2, 0); -} else { - x_138 = x_137; } -lean_ctor_set(x_138, 0, x_135); -lean_ctor_set(x_138, 1, x_136); -return x_138; } } } -else +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains(lean_object* x_1) { +_start: { -uint8_t x_139; -lean_dec(x_21); -lean_free_object(x_17); -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_4); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_MapDeclarationExtension_contains___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; lean_object* x_7; +x_6 = l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(x_1, x_2, x_3, x_4, x_5); lean_dec(x_2); lean_dec(x_1); -x_139 = !lean_is_exclusive(x_22); -if (x_139 == 0) -{ -return x_22; +x_7 = lean_box(x_6); +return x_7; } -else +} +LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_22, 0); -x_141 = lean_ctor_get(x_22, 1); -lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_22); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -return x_142; -} +uint8_t x_5; lean_object* x_6; +x_5 = l_Lean_MapDeclarationExtension_contains___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +x_6 = lean_box(x_5); +return x_6; } } -else -{ -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_143 = lean_ctor_get(x_17, 0); -x_144 = lean_ctor_get(x_17, 1); -lean_inc(x_144); -lean_inc(x_143); -lean_dec(x_17); -x_145 = lean_array_get_size(x_143); -lean_dec(x_143); -x_146 = lean_get_num_attributes(x_144); -if (lean_obj_tag(x_146) == 0) +LEAN_EXPORT lean_object* l_Lean_saveModuleData___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_147 = lean_ctor_get(x_146, 0); -lean_inc(x_147); -x_148 = lean_ctor_get(x_146, 1); -lean_inc(x_148); -lean_dec(x_146); -x_149 = lean_ctor_get(x_13, 2); -lean_inc(x_149); -lean_dec(x_13); -x_150 = lean_ctor_get(x_16, 0); -lean_inc(x_150); -if (lean_is_exclusive(x_16)) { - lean_ctor_release(x_16, 0); - lean_ctor_release(x_16, 1); - x_151 = x_16; -} else { - lean_dec_ref(x_16); - x_151 = lean_box(0); +lean_object* x_5; +x_5 = lean_save_module_data(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; } -lean_inc(x_2); -lean_inc(x_4); -x_152 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_152, 0, x_4); -lean_ctor_set(x_152, 1, x_2); -lean_inc(x_150); -x_153 = lean_apply_3(x_149, x_150, x_152, x_148); -if (lean_obj_tag(x_153) == 0) -{ -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_153, 1); -lean_inc(x_155); -lean_dec(x_153); -if (lean_is_scalar(x_151)) { - x_156 = lean_alloc_ctor(0, 2, 0); -} else { - x_156 = x_151; } -lean_ctor_set(x_156, 0, x_150); -lean_ctor_set(x_156, 1, x_154); -x_157 = l_Lean_EnvExtension_setState___rarg(x_14, x_4, x_156); -lean_dec(x_14); -x_158 = l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(x_157, x_155); -if (lean_obj_tag(x_158) == 0) +LEAN_EXPORT lean_object* l_Lean_readModuleData___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_159 = lean_ctor_get(x_158, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_158, 1); -lean_inc(x_160); -lean_dec(x_158); -x_161 = lean_st_ref_get(x_6, x_160); -x_162 = lean_ctor_get(x_161, 0); -lean_inc(x_162); -x_163 = lean_ctor_get(x_161, 1); -lean_inc(x_163); -lean_dec(x_161); -x_164 = lean_get_num_attributes(x_163); -if (lean_obj_tag(x_164) == 0) +lean_object* x_3; +x_3 = lean_read_module_data(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_164, 1); -lean_inc(x_166); -lean_dec(x_164); -x_167 = lean_array_get_size(x_162); -lean_dec(x_162); -x_168 = lean_nat_dec_lt(x_145, x_167); -lean_dec(x_167); -if (x_168 == 0) +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_2, x_3); +if (x_6 == 0) { -uint8_t x_169; -x_169 = lean_nat_dec_lt(x_147, x_165); -lean_dec(x_165); -lean_dec(x_147); -if (x_169 == 0) +lean_object* x_7; size_t x_8; lean_object* x_9; +lean_dec(x_4); +x_7 = lean_array_uget(x_1, x_2); +x_8 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_9 = lean_compacted_region_free(x_8, x_5); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_170; lean_object* x_171; -lean_dec(x_145); -x_170 = lean_box(0); -x_171 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_159, x_170, x_166); -return x_171; +lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = 1; +x_13 = lean_usize_add(x_2, x_12); +x_2 = x_13; +x_4 = x_10; +x_5 = x_11; +goto _start; } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -lean_inc(x_1); -x_172 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_159, x_1, x_145, x_166); -x_173 = lean_ctor_get(x_172, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_172, 1); -lean_inc(x_174); -lean_dec(x_172); -x_175 = lean_update_env_attributes(x_173, x_174); -if (lean_obj_tag(x_175) == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_176 = lean_ctor_get(x_175, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_175, 1); -lean_inc(x_177); -lean_dec(x_175); -x_178 = lean_box(0); -x_179 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_176, x_178, x_177); -return x_179; +return x_9; } else { -lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; -lean_dec(x_2); -lean_dec(x_1); -x_180 = lean_ctor_get(x_175, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_175, 1); -lean_inc(x_181); -if (lean_is_exclusive(x_175)) { - lean_ctor_release(x_175, 0); - lean_ctor_release(x_175, 1); - x_182 = x_175; -} else { - lean_dec_ref(x_175); - x_182 = lean_box(0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } -if (lean_is_scalar(x_182)) { - x_183 = lean_alloc_ctor(1, 2, 0); -} else { - x_183 = x_182; } -lean_ctor_set(x_183, 0, x_180); -lean_ctor_set(x_183, 1, x_181); -return x_183; +} +else +{ +lean_object* x_19; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_4); +lean_ctor_set(x_19, 1, x_5); +return x_19; } } } -else +LEAN_EXPORT lean_object* lean_environment_free_regions(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -lean_dec(x_165); -lean_dec(x_147); -lean_inc(x_1); -x_184 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_159, x_1, x_145, x_166); -x_185 = lean_ctor_get(x_184, 0); -lean_inc(x_185); -x_186 = lean_ctor_get(x_184, 1); -lean_inc(x_186); -lean_dec(x_184); -x_187 = lean_update_env_attributes(x_185, x_186); -if (lean_obj_tag(x_187) == 0) +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 5); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 2); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) { -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_187, 1); -lean_inc(x_189); -lean_dec(x_187); -x_190 = lean_box(0); -x_191 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_188, x_190, x_189); -return x_191; +lean_object* x_8; lean_object* x_9; +lean_dec(x_5); +lean_dec(x_4); +x_8 = lean_box(0); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_2); +return x_9; } else { -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -lean_dec(x_2); -lean_dec(x_1); -x_192 = lean_ctor_get(x_187, 0); -lean_inc(x_192); -x_193 = lean_ctor_get(x_187, 1); -lean_inc(x_193); -if (lean_is_exclusive(x_187)) { - lean_ctor_release(x_187, 0); - lean_ctor_release(x_187, 1); - x_194 = x_187; -} else { - lean_dec_ref(x_187); - x_194 = lean_box(0); +uint8_t x_10; +x_10 = lean_nat_dec_le(x_5, x_5); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +lean_dec(x_4); +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +return x_12; } -if (lean_is_scalar(x_194)) { - x_195 = lean_alloc_ctor(1, 2, 0); -} else { - x_195 = x_194; +else +{ +size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; +x_13 = 0; +x_14 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_15 = lean_box(0); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1(x_4, x_13, x_14, x_15, x_2); +lean_dec(x_4); +return x_16; } -lean_ctor_set(x_195, 0, x_192); -lean_ctor_set(x_195, 1, x_193); -return x_195; } } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -lean_dec(x_162); -lean_dec(x_159); -lean_dec(x_147); -lean_dec(x_145); +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1(x_1, x_6, x_7, x_4, x_5); lean_dec(x_1); -x_196 = lean_ctor_get(x_164, 0); -lean_inc(x_196); -x_197 = lean_ctor_get(x_164, 1); -lean_inc(x_197); -if (lean_is_exclusive(x_164)) { - lean_ctor_release(x_164, 0); - lean_ctor_release(x_164, 1); - x_198 = x_164; -} else { - lean_dec_ref(x_164); - x_198 = lean_box(0); -} -if (lean_is_scalar(x_198)) { - x_199 = lean_alloc_ctor(1, 2, 0); -} else { - x_199 = x_198; +return x_8; } -lean_ctor_set(x_199, 0, x_196); -lean_ctor_set(x_199, 1, x_197); -return x_199; } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) +{ +return x_4; } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; -lean_dec(x_147); -lean_dec(x_145); -lean_dec(x_2); -lean_dec(x_1); -x_200 = lean_ctor_get(x_158, 0); -lean_inc(x_200); -x_201 = lean_ctor_get(x_158, 1); -lean_inc(x_201); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_202 = x_158; -} else { - lean_dec_ref(x_158); - x_202 = lean_box(0); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = l_Lean_instInhabitedEnvExtensionState; +x_10 = l_Lean_PersistentEnvExtension_getState___rarg(x_9, x_6, x_1); +x_11 = lean_ctor_get(x_6, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_6, 4); +lean_inc(x_12); +lean_dec(x_6); +x_13 = lean_apply_1(x_12, x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_13); +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_17 = lean_array_uset(x_8, x_3, x_14); +x_3 = x_16; +x_4 = x_17; +goto _start; } -if (lean_is_scalar(x_202)) { - x_203 = lean_alloc_ctor(1, 2, 0); -} else { - x_203 = x_202; } -lean_ctor_set(x_203, 0, x_200); -lean_ctor_set(x_203, 1, x_201); -return x_203; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +switch (lean_obj_tag(x_7)) { +case 0: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_7, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_7, 1); +lean_inc(x_11); +lean_dec(x_7); +lean_inc(x_1); +x_12 = lean_apply_3(x_1, x_5, x_10, x_11); +x_3 = x_9; +x_5 = x_12; +goto _start; } -else +case 1: { -lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; -lean_dec(x_151); -lean_dec(x_150); -lean_dec(x_147); -lean_dec(x_145); +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_7, 0); +lean_inc(x_14); +lean_dec(x_7); +lean_inc(x_1); +x_15 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_1, x_14, x_5); lean_dec(x_14); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_204 = lean_ctor_get(x_153, 0); -lean_inc(x_204); -x_205 = lean_ctor_get(x_153, 1); -lean_inc(x_205); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_206 = x_153; -} else { - lean_dec_ref(x_153); - x_206 = lean_box(0); +x_3 = x_9; +x_5 = x_15; +goto _start; } -if (lean_is_scalar(x_206)) { - x_207 = lean_alloc_ctor(1, 2, 0); -} else { - x_207 = x_206; +default: +{ +x_3 = x_9; +goto _start; } -lean_ctor_set(x_207, 0, x_204); -lean_ctor_set(x_207, 1, x_205); -return x_207; } } else { -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; -lean_dec(x_145); -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_4); -lean_dec(x_2); lean_dec(x_1); -x_208 = lean_ctor_get(x_146, 0); -lean_inc(x_208); -x_209 = lean_ctor_get(x_146, 1); -lean_inc(x_209); -if (lean_is_exclusive(x_146)) { - lean_ctor_release(x_146, 0); - lean_ctor_release(x_146, 1); - x_210 = x_146; -} else { - lean_dec_ref(x_146); - x_210 = lean_box(0); -} -if (lean_is_scalar(x_210)) { - x_211 = lean_alloc_ctor(1, 2, 0); -} else { - x_211 = x_210; +return x_5; } -lean_ctor_set(x_211, 0, x_208); -lean_ctor_set(x_211, 1, x_209); -return x_211; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg___boxed), 5, 0); +return x_4; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; -x_212 = lean_ctor_get(x_7, 0); -x_213 = lean_ctor_get(x_7, 1); -lean_inc(x_213); -lean_inc(x_212); +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); lean_dec(x_7); -x_214 = lean_array_get_size(x_212); -x_215 = lean_nat_dec_lt(x_3, x_214); -lean_dec(x_214); -if (x_215 == 0) +if (x_8 == 0) { -lean_object* x_216; -lean_dec(x_212); -lean_dec(x_2); +lean_dec(x_5); lean_dec(x_1); -x_216 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_216, 0, x_4); -lean_ctor_set(x_216, 1, x_213); -return x_216; +return x_6; } else { -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; -x_217 = lean_array_fget(x_212, x_3); -lean_dec(x_212); -x_218 = lean_ctor_get(x_217, 0); -lean_inc(x_218); -x_219 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1; -x_220 = l_Lean_EnvExtension_getState___rarg(x_219, x_218, x_4); -x_221 = lean_st_ref_get(x_6, x_213); -x_222 = lean_ctor_get(x_221, 0); -lean_inc(x_222); -x_223 = lean_ctor_get(x_221, 1); -lean_inc(x_223); -if (lean_is_exclusive(x_221)) { - lean_ctor_release(x_221, 0); - lean_ctor_release(x_221, 1); - x_224 = x_221; -} else { - lean_dec_ref(x_221); - x_224 = lean_box(0); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +lean_inc(x_1); +x_11 = lean_apply_3(x_1, x_6, x_9, x_10); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_5, x_12); +lean_dec(x_5); +x_4 = lean_box(0); +x_5 = x_13; +x_6 = x_11; +goto _start; } -x_225 = lean_array_get_size(x_222); -lean_dec(x_222); -x_226 = lean_get_num_attributes(x_223); -if (lean_obj_tag(x_226) == 0) -{ -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_226, 1); -lean_inc(x_228); -lean_dec(x_226); -x_229 = lean_ctor_get(x_217, 2); -lean_inc(x_229); -lean_dec(x_217); -x_230 = lean_ctor_get(x_220, 0); -lean_inc(x_230); -if (lean_is_exclusive(x_220)) { - lean_ctor_release(x_220, 0); - lean_ctor_release(x_220, 1); - x_231 = x_220; -} else { - lean_dec_ref(x_220); - x_231 = lean_box(0); } -lean_inc(x_2); -lean_inc(x_4); -if (lean_is_scalar(x_224)) { - x_232 = lean_alloc_ctor(0, 2, 0); -} else { - x_232 = x_224; } -lean_ctor_set(x_232, 0, x_4); -lean_ctor_set(x_232, 1, x_2); -lean_inc(x_230); -x_233 = lean_apply_3(x_229, x_230, x_232, x_228); -if (lean_obj_tag(x_233) == 0) +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -x_234 = lean_ctor_get(x_233, 0); -lean_inc(x_234); -x_235 = lean_ctor_get(x_233, 1); -lean_inc(x_235); -lean_dec(x_233); -if (lean_is_scalar(x_231)) { - x_236 = lean_alloc_ctor(0, 2, 0); -} else { - x_236 = x_231; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg___boxed), 6, 0); +return x_4; } -lean_ctor_set(x_236, 0, x_230); -lean_ctor_set(x_236, 1, x_234); -x_237 = l_Lean_EnvExtension_setState___rarg(x_218, x_4, x_236); -lean_dec(x_218); -x_238 = l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(x_237, x_235); -if (lean_obj_tag(x_238) == 0) -{ -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; -x_239 = lean_ctor_get(x_238, 0); -lean_inc(x_239); -x_240 = lean_ctor_get(x_238, 1); -lean_inc(x_240); -lean_dec(x_238); -x_241 = lean_st_ref_get(x_6, x_240); -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); -x_244 = lean_get_num_attributes(x_243); -if (lean_obj_tag(x_244) == 0) +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; -x_245 = lean_ctor_get(x_244, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_244, 1); -lean_inc(x_246); -lean_dec(x_244); -x_247 = lean_array_get_size(x_242); -lean_dec(x_242); -x_248 = lean_nat_dec_lt(x_225, x_247); -lean_dec(x_247); -if (x_248 == 0) +if (lean_obj_tag(x_2) == 0) { -uint8_t x_249; -x_249 = lean_nat_dec_lt(x_227, x_245); -lean_dec(x_245); -lean_dec(x_227); -if (x_249 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) { -lean_object* x_250; lean_object* x_251; -lean_dec(x_225); -x_250 = lean_box(0); -x_251 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_239, x_250, x_246); -return x_251; +lean_dec(x_5); +lean_dec(x_1); +return x_3; } else { -lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; -lean_inc(x_1); -x_252 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_239, x_1, x_225, x_246); -x_253 = lean_ctor_get(x_252, 0); -lean_inc(x_253); -x_254 = lean_ctor_get(x_252, 1); -lean_inc(x_254); -lean_dec(x_252); -x_255 = lean_update_env_attributes(x_253, x_254); -if (lean_obj_tag(x_255) == 0) +uint8_t x_8; +x_8 = lean_nat_dec_le(x_5, x_5); +if (x_8 == 0) { -lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; -x_256 = lean_ctor_get(x_255, 0); -lean_inc(x_256); -x_257 = lean_ctor_get(x_255, 1); -lean_inc(x_257); -lean_dec(x_255); -x_258 = lean_box(0); -x_259 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_256, x_258, x_257); -return x_259; +lean_dec(x_5); +lean_dec(x_1); +return x_3; } else { -lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; -lean_dec(x_2); -lean_dec(x_1); -x_260 = lean_ctor_get(x_255, 0); -lean_inc(x_260); -x_261 = lean_ctor_get(x_255, 1); -lean_inc(x_261); -if (lean_is_exclusive(x_255)) { - lean_ctor_release(x_255, 0); - lean_ctor_release(x_255, 1); - x_262 = x_255; -} else { - lean_dec_ref(x_255); - x_262 = lean_box(0); -} -if (lean_is_scalar(x_262)) { - x_263 = lean_alloc_ctor(1, 2, 0); -} else { - x_263 = x_262; -} -lean_ctor_set(x_263, 0, x_260); -lean_ctor_set(x_263, 1, x_261); -return x_263; +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = 0; +x_10 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_11 = l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg(x_1, x_4, x_9, x_10, x_3); +return x_11; } } } else { -lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; -lean_dec(x_245); -lean_dec(x_227); -lean_inc(x_1); -x_264 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_239, x_1, x_225, x_246); -x_265 = lean_ctor_get(x_264, 0); -lean_inc(x_265); -x_266 = lean_ctor_get(x_264, 1); -lean_inc(x_266); -lean_dec(x_264); -x_267 = lean_update_env_attributes(x_265, x_266); -if (lean_obj_tag(x_267) == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_2, 0); +x_13 = lean_ctor_get(x_2, 1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg(x_1, x_12, x_13, lean_box(0), x_14, x_3); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -x_268 = lean_ctor_get(x_267, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_267, 1); -lean_inc(x_269); -lean_dec(x_267); -x_270 = lean_box(0); -x_271 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_268, x_270, x_269); -return x_271; +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg___boxed), 3, 0); +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -lean_dec(x_2); -lean_dec(x_1); -x_272 = lean_ctor_get(x_267, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_267, 1); -lean_inc(x_273); -if (lean_is_exclusive(x_267)) { - lean_ctor_release(x_267, 0); - lean_ctor_release(x_267, 1); - x_274 = x_267; -} else { - lean_dec_ref(x_267); - x_274 = lean_box(0); +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_2, x_1, x_3); +return x_4; } -if (lean_is_scalar(x_274)) { - x_275 = lean_alloc_ctor(1, 2, 0); -} else { - x_275 = x_274; } -lean_ctor_set(x_275, 0, x_272); -lean_ctor_set(x_275, 1, x_273); -return x_275; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_2, x_1, x_3); +return x_4; } } +LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_mkModuleData___spec__8(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; } else { -lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; -lean_dec(x_242); -lean_dec(x_239); -lean_dec(x_227); -lean_dec(x_225); +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 3); +lean_inc(x_5); lean_dec(x_2); -lean_dec(x_1); -x_276 = lean_ctor_get(x_244, 0); -lean_inc(x_276); -x_277 = lean_ctor_get(x_244, 1); -lean_inc(x_277); -if (lean_is_exclusive(x_244)) { - lean_ctor_release(x_244, 0); - lean_ctor_release(x_244, 1); - x_278 = x_244; -} else { - lean_dec_ref(x_244); - x_278 = lean_box(0); +x_6 = l_Lean_RBNode_fold___at_Lean_mkModuleData___spec__8(x_1, x_3); +x_7 = lean_array_push(x_6, x_4); +x_1 = x_7; +x_2 = x_5; +goto _start; } -if (lean_is_scalar(x_278)) { - x_279 = lean_alloc_ctor(1, 2, 0); -} else { - x_279 = x_278; } -lean_ctor_set(x_279, 0, x_276); -lean_ctor_set(x_279, 1, x_277); -return x_279; +} +LEAN_EXPORT lean_object* l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_3 = l_Lean_RBNode_fold___at_Lean_mkModuleData___spec__8(x_2, x_1); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -lean_dec(x_227); -lean_dec(x_225); -lean_dec(x_2); -lean_dec(x_1); -x_280 = lean_ctor_get(x_238, 0); -lean_inc(x_280); -x_281 = lean_ctor_get(x_238, 1); -lean_inc(x_281); -if (lean_is_exclusive(x_238)) { - lean_ctor_release(x_238, 0); - lean_ctor_release(x_238, 1); - x_282 = x_238; -} else { - lean_dec_ref(x_238); - x_282 = lean_box(0); +lean_object* x_4; +x_4 = lean_array_push(x_1, x_2); +return x_4; } -if (lean_is_scalar(x_282)) { - x_283 = lean_alloc_ctor(1, 2, 0); -} else { - x_283 = x_282; } -lean_ctor_set(x_283, 0, x_280); -lean_ctor_set(x_283, 1, x_281); -return x_283; +LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_array_push(x_1, x_3); +return x_4; } } -else +static lean_object* _init_l_Lean_mkModuleData___closed__1() { +_start: { -lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; -lean_dec(x_231); -lean_dec(x_230); -lean_dec(x_227); -lean_dec(x_225); -lean_dec(x_218); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_284 = lean_ctor_get(x_233, 0); -lean_inc(x_284); -x_285 = lean_ctor_get(x_233, 1); -lean_inc(x_285); -if (lean_is_exclusive(x_233)) { - lean_ctor_release(x_233, 0); - lean_ctor_release(x_233, 1); - x_286 = x_233; -} else { - lean_dec_ref(x_233); - x_286 = lean_box(0); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkModuleData___lambda__1___boxed), 3, 0); +return x_1; } -if (lean_is_scalar(x_286)) { - x_287 = lean_alloc_ctor(1, 2, 0); -} else { - x_287 = x_286; } -lean_ctor_set(x_287, 0, x_284); -lean_ctor_set(x_287, 1, x_285); -return x_287; +static lean_object* _init_l_Lean_mkModuleData___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_mkModuleData___lambda__2___boxed), 3, 0); +return x_1; +} } +LEAN_EXPORT lean_object* l_Lean_mkModuleData(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; +x_4 = lean_st_ref_get(x_3, x_2); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +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; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_array_size(x_6); +x_8 = 0; +x_9 = l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(x_1, x_7, x_8, x_6); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_mkModuleData___closed__1; +x_13 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_14 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_12, x_11, x_13); +x_15 = l_Lean_mkModuleData___closed__2; +x_16 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_15, x_11, x_13); +lean_dec(x_11); +x_17 = lean_ctor_get(x_1, 5); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_ctor_get(x_1, 4); +lean_inc(x_19); +lean_dec(x_1); +x_20 = l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(x_19); +x_21 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_14); +lean_ctor_set(x_21, 2, x_16); +lean_ctor_set(x_21, 3, x_20); +lean_ctor_set(x_21, 4, x_9); +lean_ctor_set(x_4, 0, x_21); +return x_4; } else { -lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; -lean_dec(x_225); -lean_dec(x_224); -lean_dec(x_220); -lean_dec(x_218); -lean_dec(x_217); +lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_22 = lean_ctor_get(x_4, 0); +x_23 = lean_ctor_get(x_4, 1); +lean_inc(x_23); +lean_inc(x_22); lean_dec(x_4); -lean_dec(x_2); +x_24 = lean_array_size(x_22); +x_25 = 0; +x_26 = l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(x_1, x_24, x_25, x_22); +x_27 = lean_ctor_get(x_1, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = l_Lean_mkModuleData___closed__1; +x_30 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_31 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_29, x_28, x_30); +x_32 = l_Lean_mkModuleData___closed__2; +x_33 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_32, x_28, x_30); +lean_dec(x_28); +x_34 = lean_ctor_get(x_1, 5); +lean_inc(x_34); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_ctor_get(x_1, 4); +lean_inc(x_36); lean_dec(x_1); -x_288 = lean_ctor_get(x_226, 0); -lean_inc(x_288); -x_289 = lean_ctor_get(x_226, 1); -lean_inc(x_289); -if (lean_is_exclusive(x_226)) { - lean_ctor_release(x_226, 0); - lean_ctor_release(x_226, 1); - x_290 = x_226; -} else { - lean_dec_ref(x_226); - x_290 = lean_box(0); -} -if (lean_is_scalar(x_290)) { - x_291 = lean_alloc_ctor(1, 2, 0); -} else { - x_291 = x_290; -} -lean_ctor_set(x_291, 0, x_288); -lean_ctor_set(x_291, 1, x_289); -return x_291; -} -} +x_37 = l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(x_36); +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_31); +lean_ctor_set(x_38, 2, x_33); +lean_ctor_set(x_38, 3, x_37); +lean_ctor_set(x_38, 4, x_26); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_23); +return x_39; } } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_7; -x_7 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at_Lean_mkModuleData___spec__1(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_6; -x_6 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(x_1, x_2, x_3, x_4, x_5); +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -return x_6; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_5; lean_object* x_6; -x_5 = lean_unsigned_to_nat(0u); -x_6 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(x_2, x_3, x_5, x_1, x_4); -return x_6; +lean_object* x_7; +x_7 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_mkModuleData___spec__5___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_7; } } -static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Std.Data.DHashMap.Internal.AssocList.Basic", 42, 42); -return x_1; +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; } } -static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__2() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Std.DHashMap.Internal.AssocList.get!", 36, 36); -return x_1; +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__2(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__3() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("key is not present in hash table", 32, 32); -return x_1; +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__4() { +LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__1; -x_2 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__2; -x_3 = lean_unsigned_to_nat(138u); -x_4 = lean_unsigned_to_nat(11u); -x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_4; +x_4 = l_Lean_mkModuleData___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_mkModuleData___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_3) == 0) +lean_object* x_4; +x_4 = l_Lean_mkModuleData___lambda__2(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* lean_write_module(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_4; lean_object* x_5; -x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__4; -x_5 = l_panic___rarg(x_1, x_4); -return x_5; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_inc(x_1); +x_4 = l_Lean_mkModuleData(x_1, x_3); +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_dec(x_4); +x_7 = l_Lean_Environment_mainModule(x_1); +lean_dec(x_1); +x_8 = lean_save_module_data(x_2, x_7, x_5, x_6); +lean_dec(x_5); +lean_dec(x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 1); -x_8 = lean_ctor_get(x_3, 2); -x_9 = lean_name_eq(x_6, x_2); -if (x_9 == 0) +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_name_eq(x_4, x_1); +if (x_6 == 0) { -x_3 = x_8; +x_2 = x_5; goto _start; } else { -lean_dec(x_1); -lean_inc(x_7); -return x_7; +uint8_t x_8; +x_8 = 1; +return x_8; } } } } -static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__1() { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_mkExtNameMap___spec__4(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("import ", 7, 7); +if (lean_obj_tag(x_2) == 0) +{ return x_1; } +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = l_Lean_Name_hash___override(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = l_Lean_Name_hash___override(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} } -static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__2() { +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_mkExtNameMap___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" failed, environment already contains '", 39, 39); -return x_1; +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_mkExtNameMap___spec__4(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; } } -static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__3() { +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("' from ", 7, 7); -return x_1; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_mkExtNameMap___spec__3(x_7, x_1, x_6); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_array_get_size(x_6); -x_8 = lean_nat_dec_lt(x_3, x_7); -x_9 = 1; -x_10 = l_Lean_instToStringImport___closed__1; -lean_inc(x_4); -x_11 = l_Lean_Name_toString(x_4, x_9, x_10); -if (x_8 == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_88; lean_object* x_89; -x_88 = l_Lean_instInhabitedName; -x_89 = l_outOfBounds___rarg(x_88); -x_12 = x_89; -goto block_87; +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; } else { -lean_object* x_90; -x_90 = lean_array_fget(x_6, x_3); -x_12 = x_90; -goto block_87; -} -block_87: +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_13 = l_Lean_Name_toString(x_12, x_9, x_10); -x_14 = l_Lean_throwAlreadyImported___rarg___closed__1; -x_15 = lean_string_append(x_14, x_13); -lean_dec(x_13); -x_16 = l_Lean_throwAlreadyImported___rarg___closed__2; -x_17 = lean_string_append(x_15, x_16); -x_18 = lean_string_append(x_17, x_11); -lean_dec(x_11); -x_19 = l_Lean_throwAlreadyImported___rarg___closed__3; -x_20 = lean_string_append(x_18, x_19); -x_21 = !lean_is_exclusive(x_2); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; size_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_22 = lean_ctor_get(x_2, 1); -x_23 = lean_ctor_get(x_2, 0); -lean_dec(x_23); -x_24 = lean_array_get_size(x_22); -x_25 = l_Lean_Name_hash___override(x_4); -x_26 = 32; -x_27 = lean_uint64_shift_right(x_25, x_26); -x_28 = lean_uint64_xor(x_25, x_27); -x_29 = 16; -x_30 = lean_uint64_shift_right(x_28, x_29); -x_31 = lean_uint64_xor(x_28, x_30); -x_32 = lean_uint64_to_usize(x_31); -x_33 = lean_usize_of_nat(x_24); -lean_dec(x_24); -x_34 = 1; -x_35 = lean_usize_sub(x_33, x_34); -x_36 = lean_usize_land(x_32, x_35); -x_37 = lean_array_uget(x_22, x_36); -lean_dec(x_22); -x_38 = l_Lean_instInhabitedModuleIdx; -x_39 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(x_38, x_4, x_37); -lean_dec(x_37); -lean_dec(x_4); -x_40 = lean_nat_dec_lt(x_39, x_7); -lean_dec(x_7); -if (x_40 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_name_eq(x_6, x_1); +if (x_9 == 0) { -lean_object* x_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_dec(x_39); -x_41 = l_Lean_instInhabitedName; -x_42 = l_outOfBounds___rarg(x_41); -x_43 = l_Lean_Name_toString(x_42, x_9, x_10); -x_44 = lean_string_append(x_20, x_43); -lean_dec(x_43); -x_45 = l_Lean_instToStringImport___closed__2; -x_46 = lean_string_append(x_44, x_45); -x_47 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set_tag(x_2, 1); -lean_ctor_set(x_2, 1, x_5); -lean_ctor_set(x_2, 0, x_47); -return x_2; +lean_object* x_10; +x_10 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_10); +return x_3; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_48 = lean_array_fget(x_6, x_39); -lean_dec(x_39); -x_49 = l_Lean_Name_toString(x_48, x_9, x_10); -x_50 = lean_string_append(x_20, x_49); -lean_dec(x_49); -x_51 = l_Lean_instToStringImport___closed__2; -x_52 = lean_string_append(x_50, x_51); -x_53 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set_tag(x_2, 1); -lean_ctor_set(x_2, 1, x_5); -lean_ctor_set(x_2, 0, x_53); -return x_2; +lean_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } else { -lean_object* x_54; lean_object* x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; size_t x_63; size_t x_64; size_t x_65; size_t x_66; size_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_54 = lean_ctor_get(x_2, 1); -lean_inc(x_54); -lean_dec(x_2); -x_55 = lean_array_get_size(x_54); -x_56 = l_Lean_Name_hash___override(x_4); -x_57 = 32; -x_58 = lean_uint64_shift_right(x_56, x_57); -x_59 = lean_uint64_xor(x_56, x_58); -x_60 = 16; -x_61 = lean_uint64_shift_right(x_59, x_60); -x_62 = lean_uint64_xor(x_59, x_61); -x_63 = lean_uint64_to_usize(x_62); -x_64 = lean_usize_of_nat(x_55); -lean_dec(x_55); -x_65 = 1; -x_66 = lean_usize_sub(x_64, x_65); -x_67 = lean_usize_land(x_63, x_66); -x_68 = lean_array_uget(x_54, x_67); -lean_dec(x_54); -x_69 = l_Lean_instInhabitedModuleIdx; -x_70 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(x_69, x_4, x_68); -lean_dec(x_68); -lean_dec(x_4); -x_71 = lean_nat_dec_lt(x_70, x_7); -lean_dec(x_7); -if (x_71 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_3, 2); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +x_14 = lean_name_eq(x_11, x_1); +if (x_14 == 0) { -lean_object* x_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_dec(x_70); -x_72 = l_Lean_instInhabitedName; -x_73 = l_outOfBounds___rarg(x_72); -x_74 = l_Lean_Name_toString(x_73, x_9, x_10); -x_75 = lean_string_append(x_20, x_74); -lean_dec(x_74); -x_76 = l_Lean_instToStringImport___closed__2; -x_77 = lean_string_append(x_75, x_76); -x_78 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_78, 0, x_77); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_5); -return x_79; +lean_object* x_15; lean_object* x_16; +x_15 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_1, x_2, x_13); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_15); +return x_16; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_80 = lean_array_fget(x_6, x_70); -lean_dec(x_70); -x_81 = l_Lean_Name_toString(x_80, x_9, x_10); -x_82 = lean_string_append(x_20, x_81); -lean_dec(x_81); -x_83 = l_Lean_instToStringImport___closed__2; -x_84 = lean_string_append(x_82, x_83); -x_85 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_85, 0, x_84); -x_86 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_5); -return x_86; +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 2, x_13); +return x_17; } } } } } -LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___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) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_throwAlreadyImported___rarg___boxed), 5, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_10; uint8_t x_11; +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_nat_dec_lt(x_6, x_10); +if (x_11 == 0) { -lean_object* x_4; -x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(x_1, x_2, x_3); -lean_dec(x_3); +lean_object* x_12; +lean_dec(x_6); lean_dec(x_2); -return x_4; -} +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_9); +return x_12; } -LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; -x_6 = l_Lean_throwAlreadyImported___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -lean_dec(x_1); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_ImportStateM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_array_fget(x_1, x_6); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = !lean_is_exclusive(x_5); +if (x_15 == 0) { -lean_object* x_4; uint8_t x_5; -x_4 = lean_st_mk_ref(x_2, x_3); -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; lean_object* x_31; uint8_t x_32; +x_16 = lean_ctor_get(x_5, 0); +x_17 = lean_ctor_get(x_5, 1); +x_18 = lean_array_get_size(x_17); +x_19 = l_Lean_Name_hash___override(x_14); +x_20 = 32; +x_21 = lean_uint64_shift_right(x_19, x_20); +x_22 = lean_uint64_xor(x_19, x_21); +x_23 = 16; +x_24 = lean_uint64_shift_right(x_22, x_23); +x_25 = lean_uint64_xor(x_22, x_24); +x_26 = lean_uint64_to_usize(x_25); +x_27 = lean_usize_of_nat(x_18); +lean_dec(x_18); +x_28 = 1; +x_29 = lean_usize_sub(x_27, x_28); +x_30 = lean_usize_land(x_26, x_29); +x_31 = lean_array_uget(x_17, x_30); +x_32 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_14, x_31); +if (x_32 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_4, 0); -x_7 = lean_ctor_get(x_4, 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; uint8_t x_42; +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_add(x_16, x_33); +lean_dec(x_16); lean_inc(x_6); -x_8 = lean_apply_2(x_1, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_14); +lean_ctor_set(x_35, 1, x_6); +lean_ctor_set(x_35, 2, x_31); +x_36 = lean_array_uset(x_17, x_30, x_35); +x_37 = lean_unsigned_to_nat(4u); +x_38 = lean_nat_mul(x_34, x_37); +x_39 = lean_unsigned_to_nat(3u); +x_40 = lean_nat_div(x_38, x_39); +lean_dec(x_38); +x_41 = lean_array_get_size(x_36); +x_42 = lean_nat_dec_le(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +if (x_42 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_st_ref_get(x_6, x_10); -lean_dec(x_6); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; -x_13 = lean_ctor_get(x_11, 0); -lean_ctor_set(x_4, 1, x_13); -lean_ctor_set(x_4, 0, x_9); -lean_ctor_set(x_11, 0, x_4); -return x_11; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(x_36); +lean_ctor_set(x_5, 1, x_43); +lean_ctor_set(x_5, 0, x_34); +x_44 = lean_ctor_get(x_4, 2); +x_45 = lean_nat_add(x_6, x_44); +lean_dec(x_6); +x_6 = x_45; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 0); -x_15 = lean_ctor_get(x_11, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_11); -lean_ctor_set(x_4, 1, x_14); -lean_ctor_set(x_4, 0, x_9); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_4); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_object* x_47; lean_object* x_48; +lean_ctor_set(x_5, 1, x_36); +lean_ctor_set(x_5, 0, x_34); +x_47 = lean_ctor_get(x_4, 2); +x_48 = lean_nat_add(x_6, x_47); +lean_dec(x_6); +x_6 = x_48; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; } } else { -uint8_t x_17; -lean_free_object(x_4); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_inc(x_2); +x_50 = lean_array_uset(x_17, x_30, x_2); +lean_inc(x_6); +x_51 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_14, x_6, x_31); +x_52 = lean_array_uset(x_50, x_30, x_51); +lean_ctor_set(x_5, 1, x_52); +x_53 = lean_ctor_get(x_4, 2); +x_54 = lean_nat_add(x_6, x_53); lean_dec(x_6); -x_17 = !lean_is_exclusive(x_8); -if (x_17 == 0) -{ -return x_8; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_8, 0); -x_19 = lean_ctor_get(x_8, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_8); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; -} +x_6 = x_54; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; } } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_4, 0); -x_22 = lean_ctor_get(x_4, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_4); -lean_inc(x_21); -x_23 = lean_apply_2(x_1, x_21, x_22); -if (lean_obj_tag(x_23) == 0) +lean_object* x_56; lean_object* x_57; lean_object* x_58; uint64_t x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; uint64_t x_65; size_t x_66; size_t x_67; size_t x_68; size_t x_69; size_t x_70; lean_object* x_71; uint8_t x_72; +x_56 = lean_ctor_get(x_5, 0); +x_57 = lean_ctor_get(x_5, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_5); +x_58 = lean_array_get_size(x_57); +x_59 = l_Lean_Name_hash___override(x_14); +x_60 = 32; +x_61 = lean_uint64_shift_right(x_59, x_60); +x_62 = lean_uint64_xor(x_59, x_61); +x_63 = 16; +x_64 = lean_uint64_shift_right(x_62, x_63); +x_65 = lean_uint64_xor(x_62, x_64); +x_66 = lean_uint64_to_usize(x_65); +x_67 = lean_usize_of_nat(x_58); +lean_dec(x_58); +x_68 = 1; +x_69 = lean_usize_sub(x_67, x_68); +x_70 = lean_usize_land(x_66, x_69); +x_71 = lean_array_uget(x_57, x_70); +x_72 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_14, x_71); +if (x_72 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_st_ref_get(x_21, x_25); -lean_dec(x_21); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_29 = x_26; -} else { - lean_dec_ref(x_26); - x_29 = lean_box(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; uint8_t x_82; +x_73 = lean_unsigned_to_nat(1u); +x_74 = lean_nat_add(x_56, x_73); +lean_dec(x_56); +lean_inc(x_6); +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_14); +lean_ctor_set(x_75, 1, x_6); +lean_ctor_set(x_75, 2, x_71); +x_76 = lean_array_uset(x_57, x_70, x_75); +x_77 = lean_unsigned_to_nat(4u); +x_78 = lean_nat_mul(x_74, x_77); +x_79 = lean_unsigned_to_nat(3u); +x_80 = lean_nat_div(x_78, x_79); +lean_dec(x_78); +x_81 = lean_array_get_size(x_76); +x_82 = lean_nat_dec_le(x_80, x_81); +lean_dec(x_81); +lean_dec(x_80); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_83 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_mkExtNameMap___spec__2(x_76); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_74); +lean_ctor_set(x_84, 1, x_83); +x_85 = lean_ctor_get(x_4, 2); +x_86 = lean_nat_add(x_6, x_85); +lean_dec(x_6); +x_5 = x_84; +x_6 = x_86; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; } -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_24); -lean_ctor_set(x_30, 1, x_27); -if (lean_is_scalar(x_29)) { - x_31 = lean_alloc_ctor(0, 2, 0); -} else { - x_31 = x_29; +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_74); +lean_ctor_set(x_88, 1, x_76); +x_89 = lean_ctor_get(x_4, 2); +x_90 = lean_nat_add(x_6, x_89); +lean_dec(x_6); +x_5 = x_88; +x_6 = x_90; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; } -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_28); -return x_31; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_dec(x_21); -x_32 = lean_ctor_get(x_23, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_23, 1); -lean_inc(x_33); -if (lean_is_exclusive(x_23)) { - lean_ctor_release(x_23, 0); - lean_ctor_release(x_23, 1); - x_34 = x_23; -} else { - lean_dec_ref(x_23); - x_34 = lean_box(0); -} -if (lean_is_scalar(x_34)) { - x_35 = lean_alloc_ctor(1, 2, 0); -} else { - x_35 = x_34; +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_inc(x_2); +x_92 = lean_array_uset(x_57, x_70, x_2); +lean_inc(x_6); +x_93 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_mkExtNameMap___spec__5(x_14, x_6, x_71); +x_94 = lean_array_uset(x_92, x_70, x_93); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_56); +lean_ctor_set(x_95, 1, x_94); +x_96 = lean_ctor_get(x_4, 2); +x_97 = lean_nat_add(x_6, x_96); +lean_dec(x_6); +x_5 = x_95; +x_6 = x_97; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; } -lean_ctor_set(x_35, 0, x_32); -lean_ctor_set(x_35, 1, x_33); -return x_35; } } } } -LEAN_EXPORT lean_object* l_Lean_ImportStateM_run(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkExtNameMap(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_ImportStateM_run___rarg), 3, 0); -return x_2; -} +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_3 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; +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_dec(x_4); +x_7 = lean_box(0); +x_8 = lean_array_get_size(x_5); +x_9 = lean_unsigned_to_nat(1u); +lean_inc(x_1); +x_10 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_8); +lean_ctor_set(x_10, 2, x_9); +x_11 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; +x_12 = l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spec__6(x_5, x_7, x_10, x_10, x_11, x_1, lean_box(0), lean_box(0), x_6); +lean_dec(x_10); +lean_dec(x_5); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +return x_12; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_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) { +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_6; -x_6 = lean_read_module_data(x_1, x_5); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -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); -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 = 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); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; 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_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -x_27 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1; -lean_ctor_set(x_24, 0, x_27); -return x_24; +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_mkExtNameMap___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -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_forIn_x27Unsafe_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; +lean_object* x_10; +x_10 = l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; 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_39); - x_41 = lean_box(0); +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_array_get_size(x_1); +x_4 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_5 = lean_mk_array(x_3, x_4); +x_6 = !lean_is_exclusive(x_2); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_2, 0); +lean_dec(x_7); +lean_ctor_set(x_2, 0, x_5); +return x_2; } -x_42 = l_Array_forIn_x27Unsafe_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_43 = x_41; +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_8); +return x_9; } -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; } } -else +LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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) { +_start: { -uint8_t x_44; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_2); -x_44 = !lean_is_exclusive(x_12); -if (x_44 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_4, x_3); +if (x_7 == 0) { -return x_12; +lean_object* x_8; +lean_dec(x_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_6); +return x_8; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_12, 0); -x_46 = lean_ctor_get(x_12, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_12); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; +x_9 = lean_ctor_get(x_2, 0); +x_10 = lean_array_uget(x_9, x_4); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +lean_inc(x_1); +x_12 = lean_alloc_closure((void*)(l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1___boxed), 2, 1); +lean_closure_set(x_12, 0, x_1); +x_13 = l_Lean_EnvExtension_modifyState___rarg(x_11, x_5, x_12); +lean_dec(x_11); +x_14 = 1; +x_15 = lean_usize_add(x_4, x_14); +x_4 = x_15; +x_5 = x_13; +goto _start; +} } } +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } else { -uint8_t x_48; -lean_dec(x_2); -x_48 = !lean_is_exclusive(x_6); -if (x_48 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_name_eq(x_4, x_1); +if (x_7 == 0) { -return x_6; +x_2 = x_6; +goto _start; } else { -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; +lean_object* x_9; +lean_inc(x_5); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_5); +return x_9; } } } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1() { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("object file '", 13, 13); -return x_1; -} +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_array_set(x_5, x_1, x_2); +lean_ctor_set(x_3, 0, x_6); +return x_3; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("' of module ", 12, 12); -return x_1; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_ctor_get(x_3, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_3); +x_9 = lean_array_set(x_7, x_1, x_2); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3() { +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" does not exist", 15, 15); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_instInhabitedEnvExtensionState; +x_2 = l_Lean_instInhabitedPersistentEnvExtension(lean_box(0), lean_box(0), lean_box(0), x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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) { _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) +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_8, x_7); +if (x_11 == 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_10); -x_11 = l_Lean_NameHashSet_insert(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); -x_14 = l_Lean_findOLean(x_10, x_13); -if (lean_obj_tag(x_14) == 0) +lean_object* x_12; +lean_dec(x_3); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_9); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +else { -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_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; +x_13 = lean_array_uget(x_6, x_8); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); +lean_dec(x_13); +x_16 = lean_ctor_get(x_2, 1); +x_17 = lean_array_get_size(x_16); +x_18 = l_Lean_Name_hash___override(x_14); +x_19 = 32; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = 16; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = lean_uint64_to_usize(x_24); +x_26 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_27 = 1; +x_28 = lean_usize_sub(x_26, x_27); +x_29 = lean_usize_land(x_25, x_28); +x_30 = lean_array_uget(x_16, x_29); +x_31 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(x_14, x_30); +lean_dec(x_30); 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) -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_17); -if (x_20 == 0) +if (lean_obj_tag(x_31) == 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; lean_object* x_32; -x_21 = lean_ctor_get(x_17, 0); -lean_dec(x_21); -x_22 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; -x_23 = lean_string_append(x_22, x_15); +size_t x_32; lean_dec(x_15); -x_24 = l_Array_forIn_x27Unsafe_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_instToStringImport___closed__1; -x_28 = l_Lean_Name_toString(x_10, x_26, x_27); -x_29 = lean_string_append(x_25, x_28); -lean_dec(x_28); -x_30 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___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_17, 1); -lean_ctor_set(x_17, 0, x_32); -return x_17; +x_32 = lean_usize_add(x_8, x_27); +x_8 = x_32; +goto _start; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_33 = lean_ctor_get(x_17, 1); -lean_inc(x_33); -lean_dec(x_17); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; -x_35 = lean_string_append(x_34, x_15); -lean_dec(x_15); -x_36 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; -x_37 = lean_string_append(x_35, x_36); -x_38 = 1; -x_39 = l_Lean_instToStringImport___closed__1; -x_40 = l_Lean_Name_toString(x_10, x_38, x_39); -x_41 = lean_string_append(x_37, x_40); +lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_array_get_size(x_1); +x_36 = lean_nat_dec_lt(x_34, x_35); +lean_dec(x_35); +lean_inc(x_3); +x_37 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1___boxed), 3, 2); +lean_closure_set(x_37, 0, x_3); +lean_closure_set(x_37, 1, x_15); +if (x_36 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; +lean_dec(x_34); +x_38 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___closed__1; +x_39 = l_outOfBounds___rarg(x_38); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +lean_dec(x_39); +x_41 = l_Lean_EnvExtension_modifyState___rarg(x_40, x_9, x_37); lean_dec(x_40); -x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; -x_43 = lean_string_append(x_41, x_42); -x_44 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_33); -return x_45; -} +x_42 = lean_usize_add(x_8, x_27); +x_8 = x_42; +x_9 = x_41; +goto _start; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_17, 1); -lean_inc(x_46); -lean_dec(x_17); -x_47 = lean_box(0); -x_48 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_15, x_10, x_47, x_3, x_46); -lean_dec(x_15); -return x_48; +lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; +x_44 = lean_array_fget(x_1, x_34); +lean_dec(x_34); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +lean_dec(x_44); +x_46 = l_Lean_EnvExtension_modifyState___rarg(x_45, x_9, x_37); +lean_dec(x_45); +x_47 = lean_usize_add(x_8, x_27); +x_8 = x_47; +x_9 = x_46; +goto _start; } } -else +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_49; -lean_dec(x_10); -x_49 = !lean_is_exclusive(x_14); -if (x_49 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_5, 1); +x_12 = lean_nat_dec_lt(x_7, x_11); +if (x_12 == 0) { -return x_14; +lean_object* x_13; +lean_dec(x_7); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_10); +return x_13; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_14, 0); -x_51 = lean_ctor_get(x_14, 1); -lean_inc(x_51); -lean_inc(x_50); +lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_14 = lean_array_fget(x_1, x_7); +x_15 = lean_ctor_get(x_14, 4); +lean_inc(x_15); lean_dec(x_14); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +x_16 = lean_box(0); +x_17 = lean_array_size(x_15); +x_18 = 0; +lean_inc(x_7); +x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3(x_2, x_3, x_7, x_15, x_16, x_15, x_17, x_18, x_6, x_10); +lean_dec(x_15); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_ctor_get(x_5, 2); +x_23 = lean_nat_add(x_7, x_22); +lean_dec(x_7); +x_6 = x_20; +x_7 = x_23; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_10 = x_21; +goto _start; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_setImportedEntries(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_53 = lean_ctor_get(x_6, 0); -x_54 = lean_ctor_get(x_6, 1); -x_55 = lean_ctor_get(x_6, 2); -x_56 = lean_ctor_get(x_6, 3); -lean_inc(x_56); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_6); -x_57 = lean_ctor_get(x_1, 0); -lean_inc(x_57); -lean_dec(x_1); -lean_inc(x_57); -x_58 = l_Lean_NameHashSet_insert(x_53, x_57); -x_59 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_54); -lean_ctor_set(x_59, 2, x_55); -lean_ctor_set(x_59, 3, x_56); -x_60 = lean_st_ref_set(x_3, x_59, x_7); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_62 = l_Lean_findOLean(x_57, x_61); -if (lean_obj_tag(x_62) == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; lean_object* x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_5 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; +x_6 = lean_st_ref_get(x_5, x_4); +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_array_get_size(x_7); +lean_inc(x_3); +lean_inc(x_7); +x_10 = l_Array_toSubarray___rarg(x_7, x_3, x_9); +x_11 = lean_ctor_get(x_10, 2); +lean_inc(x_11); +x_12 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +x_14 = lean_usize_of_nat(x_13); +lean_dec(x_13); +lean_inc(x_2); +x_15 = l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1(x_2, x_10, x_12, x_14, x_1, x_8); +lean_dec(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 = l_Lean_mkExtNameMap(x_3, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_array_get_size(x_2); +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_23); +x_25 = l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4(x_2, x_7, x_19, x_24, x_24, x_16, x_22, lean_box(0), lean_box(0), x_20); +lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_7); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = l_System_FilePath_pathExists(x_63, x_64); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_unbox(x_66); -lean_dec(x_66); -if (x_67 == 0) +return x_25; +} +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; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_68 = lean_ctor_get(x_65, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_69 = x_65; -} else { - lean_dec_ref(x_65); - x_69 = lean_box(0); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_25); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } -x_70 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; -x_71 = lean_string_append(x_70, x_63); -lean_dec(x_63); -x_72 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; -x_73 = lean_string_append(x_71, x_72); -x_74 = 1; -x_75 = l_Lean_instToStringImport___closed__1; -x_76 = l_Lean_Name_toString(x_57, x_74, x_75); -x_77 = lean_string_append(x_73, x_76); -lean_dec(x_76); -x_78 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; -x_79 = lean_string_append(x_77, x_78); -x_80 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_80, 0, x_79); -if (lean_is_scalar(x_69)) { - x_81 = lean_alloc_ctor(1, 2, 0); -} else { - x_81 = x_69; - lean_ctor_set_tag(x_81, 1); } -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_68); -return x_81; } -else +LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_65, 1); -lean_inc(x_82); -lean_dec(x_65); -x_83 = lean_box(0); -x_84 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_63, x_57, x_83, x_3, x_82); -lean_dec(x_63); -return x_84; +lean_object* x_3; +x_3 = l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1(x_1, x_2); +lean_dec(x_1); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -lean_dec(x_57); -x_85 = lean_ctor_get(x_62, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_62, 1); -lean_inc(x_86); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_87 = x_62; -} else { - lean_dec_ref(x_62); - x_87 = lean_box(0); +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_9 = l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1(x_1, x_2, x_7, x_8, x_5, x_6); +lean_dec(x_2); +return x_9; } -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; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___lambda__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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: { -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_5, x_4); -if (x_9 == 0) +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_12 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_12, x_9, x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_10; -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_6); -lean_ctor_set(x_10, 1, x_8); -return x_10; +lean_object* x_11; +x_11 = l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; } -else +} +LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -lean_dec(x_6); -x_11 = lean_array_uget(x_3, x_5); -x_12 = lean_st_ref_get(x_7, x_8); -x_13 = lean_ctor_get_uint8(x_11, sizeof(void*)*1); -if (x_13 == 0) +lean_object* x_3; +x_3 = lean_update_env_attributes(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_getNumBuiltinAttributes___boxed(lean_object* x_1) { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_ctor_get(x_11, 0); -lean_inc(x_17); -x_18 = l_Lean_NameHashSet_contains(x_16, x_17); -lean_dec(x_17); -lean_dec(x_16); -if (x_18 == 0) +lean_object* x_2; +x_2 = lean_get_num_attributes(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_19; lean_object* x_20; -x_19 = lean_box(0); -x_20 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(x_11, x_19, x_7, x_15); -if (lean_obj_tag(x_20) == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_unsigned_to_nat(1u); +x_8 = lean_nat_add(x_1, x_7); +x_9 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(x_2, x_3, x_8, x_4, x_6); +lean_dec(x_8); +return x_9; +} +} +static lean_object* _init_l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1() { +_start: { -lean_object* x_21; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -if (lean_obj_tag(x_21) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_instInhabitedEnvExtensionState; +x_2 = l_Lean_instInhabitedPersistentEnvExtensionState___rarg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_20); -if (x_22 == 0) +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; +x_7 = lean_st_ref_get(x_6, x_5); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 0); -lean_dec(x_23); -x_24 = lean_ctor_get(x_21, 0); -lean_inc(x_24); -lean_dec(x_21); -lean_ctor_set(x_20, 0, x_24); -return x_20; +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_7, 1); +x_11 = lean_array_get_size(x_9); +x_12 = lean_nat_dec_lt(x_3, x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +lean_ctor_set(x_7, 0, x_4); +return x_7; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_20, 1); +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_free_object(x_7); +x_13 = lean_array_fget(x_9, x_3); +lean_dec(x_9); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1; +x_16 = l_Lean_EnvExtension_getState___rarg(x_15, x_14, x_4); +x_17 = lean_st_ref_get(x_6, x_10); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +x_21 = lean_array_get_size(x_19); +lean_dec(x_19); +x_22 = lean_get_num_attributes(x_20); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_13, 2); lean_inc(x_25); -lean_dec(x_20); -x_26 = lean_ctor_get(x_21, 0); -lean_inc(x_26); +lean_dec(x_13); +x_26 = !lean_is_exclusive(x_16); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_16, 0); +x_28 = lean_ctor_get(x_16, 1); +lean_dec(x_28); +lean_inc(x_2); +lean_inc(x_4); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 0, x_4); +lean_inc(x_27); +x_29 = lean_apply_3(x_25, x_27, x_17, x_24); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +lean_ctor_set(x_16, 1, x_30); +x_32 = l_Lean_EnvExtension_setState___rarg(x_14, x_4, x_16); +lean_dec(x_14); +x_33 = l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(x_32, x_31); +if (lean_obj_tag(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 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_st_ref_get(x_6, x_35); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_get_num_attributes(x_38); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_array_get_size(x_37); +lean_dec(x_37); +x_43 = lean_nat_dec_lt(x_21, x_42); +lean_dec(x_42); +if (x_43 == 0) +{ +uint8_t x_44; +x_44 = lean_nat_dec_lt(x_23, x_40); +lean_dec(x_40); +lean_dec(x_23); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_dec(x_21); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; -} +x_45 = lean_box(0); +x_46 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_34, x_45, x_41); +return x_46; } else { -lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; -x_28 = lean_ctor_get(x_20, 1); -lean_inc(x_28); -lean_dec(x_20); -x_29 = lean_ctor_get(x_21, 0); -lean_inc(x_29); -lean_dec(x_21); -x_30 = 1; -x_31 = lean_usize_add(x_5, x_30); -x_5 = x_31; -x_6 = x_29; -x_8 = x_28; -goto _start; -} +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_inc(x_1); +x_47 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_34, x_1, x_21, x_41); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = lean_update_env_attributes(x_48, x_49); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_box(0); +x_54 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_51, x_53, x_52); +return x_54; } else { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_20); -if (x_33 == 0) +uint8_t x_55; +lean_dec(x_2); +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_50); +if (x_55 == 0) { -return x_20; +return x_50; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_20, 0); -x_35 = lean_ctor_get(x_20, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_20); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_50, 0); +x_57 = lean_ctor_get(x_50, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_50); +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 { -size_t x_37; size_t x_38; lean_object* x_39; -lean_dec(x_11); -x_37 = 1; -x_38 = lean_usize_add(x_5, x_37); -x_39 = lean_box(0); -x_5 = x_38; -x_6 = x_39; -x_8 = x_15; -goto _start; -} +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_40); +lean_dec(x_23); +lean_inc(x_1); +x_59 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_34, x_1, x_21, x_41); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec(x_59); +x_62 = lean_update_env_attributes(x_60, x_61); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_box(0); +x_66 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_63, x_65, x_64); +return x_66; } else { -lean_object* x_41; size_t x_42; size_t x_43; lean_object* x_44; -lean_dec(x_11); -x_41 = lean_ctor_get(x_12, 1); -lean_inc(x_41); -lean_dec(x_12); -x_42 = 1; -x_43 = lean_usize_add(x_5, x_42); -x_44 = lean_box(0); -x_5 = x_43; -x_6 = x_44; -x_8 = x_41; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_importModulesCore(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; size_t x_5; size_t x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_box(0); -x_5 = lean_array_size(x_1); -x_6 = 0; -x_7 = lean_box(0); -x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1(x_1, x_4, x_1, x_5, x_6, x_7, x_2, x_3); -if (lean_obj_tag(x_8) == 0) -{ -uint8_t x_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +uint8_t x_67; +lean_dec(x_2); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_62); +if (x_67 == 0) { -lean_object* x_10; -x_10 = lean_ctor_get(x_8, 0); -lean_dec(x_10); -lean_ctor_set(x_8, 0, x_7); -return x_8; +return x_62; } else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_dec(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_7); -lean_ctor_set(x_12, 1, x_11); -return x_12; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_62, 0); +x_69 = lean_ctor_get(x_62, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_62); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} } } else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_8); -if (x_13 == 0) +uint8_t x_71; +lean_dec(x_37); +lean_dec(x_34); +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_2); +lean_dec(x_1); +x_71 = !lean_is_exclusive(x_39); +if (x_71 == 0) { -return x_8; +return x_39; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_8, 0); -x_15 = lean_ctor_get(x_8, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_8); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_39, 0); +x_73 = lean_ctor_get(x_39, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_39); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_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: +else { -lean_object* x_6; -x_6 = l_Array_forIn_x27Unsafe_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); +uint8_t x_75; +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_2); lean_dec(x_1); -return x_6; -} +x_75 = !lean_is_exclusive(x_33); +if (x_75 == 0) +{ +return x_33; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_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: +else { -lean_object* x_5; -x_5 = l_Array_forIn_x27Unsafe_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; +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_33, 0); +x_77 = lean_ctor_get(x_33, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_33); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +} +else { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_4); +uint8_t x_79; +lean_free_object(x_16); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_14); lean_dec(x_4); -x_10 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_11 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_11; -} +x_79 = !lean_is_exclusive(x_29); +if (x_79 == 0) +{ +return x_29; } -LEAN_EXPORT lean_object* l_Lean_importModulesCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_importModulesCore(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_4; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_29, 0); +x_81 = lean_ctor_get(x_29, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_29); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } -LEAN_EXPORT uint8_t l___private_Lean_Environment_0__Lean_equivInfo(lean_object* x_1, lean_object* x_2) { -_start: +} +else { -if (lean_obj_tag(x_1) == 2) +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_16, 0); +lean_inc(x_83); +lean_dec(x_16); +lean_inc(x_2); +lean_inc(x_4); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 0, x_4); +lean_inc(x_83); +x_84 = lean_apply_3(x_25, x_83, x_17, x_24); +if (lean_obj_tag(x_84) == 0) { -if (lean_obj_tag(x_2) == 2) +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_83); +lean_ctor_set(x_87, 1, x_85); +x_88 = l_Lean_EnvExtension_setState___rarg(x_14, x_4, x_87); +lean_dec(x_14); +x_89 = l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(x_88, x_86); +if (lean_obj_tag(x_89) == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_3 = lean_ctor_get(x_1, 0); -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_5, 0); -x_7 = lean_ctor_get(x_4, 0); -x_8 = lean_ctor_get(x_7, 0); -x_9 = lean_name_eq(x_6, x_8); -if (x_9 == 0) +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +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_st_ref_get(x_6, x_91); +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_get_num_attributes(x_94); +if (lean_obj_tag(x_95) == 0) { -uint8_t x_10; -x_10 = 0; -return x_10; -} -else +lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_array_get_size(x_93); +lean_dec(x_93); +x_99 = lean_nat_dec_lt(x_21, x_98); +lean_dec(x_98); +if (x_99 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_5, 2); -x_12 = lean_ctor_get(x_7, 2); -x_13 = lean_expr_eqv(x_11, x_12); -if (x_13 == 0) +uint8_t x_100; +x_100 = lean_nat_dec_lt(x_23, x_96); +lean_dec(x_96); +lean_dec(x_23); +if (x_100 == 0) { -uint8_t x_14; -x_14 = 0; -return x_14; +lean_object* x_101; lean_object* x_102; +lean_dec(x_21); +x_101 = lean_box(0); +x_102 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_90, x_101, x_97); +return x_102; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_5, 1); -x_16 = lean_ctor_get(x_7, 1); -x_17 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_431____spec__1(x_15, x_16); -if (x_17 == 0) +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_inc(x_1); +x_103 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_90, x_1, x_21, x_97); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); +x_106 = lean_update_env_attributes(x_104, x_105); +if (lean_obj_tag(x_106) == 0) { -uint8_t x_18; -x_18 = 0; -return x_18; +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_box(0); +x_110 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_107, x_109, x_108); +return x_110; } else { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_3, 2); -x_20 = lean_ctor_get(x_4, 2); -x_21 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_431____spec__1(x_19, x_20); -return x_21; +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_dec(x_2); +lean_dec(x_1); +x_111 = lean_ctor_get(x_106, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_106, 1); +lean_inc(x_112); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_113 = x_106; +} else { + lean_dec_ref(x_106); + x_113 = lean_box(0); +} +if (lean_is_scalar(x_113)) { + x_114 = lean_alloc_ctor(1, 2, 0); +} else { + x_114 = x_113; } +lean_ctor_set(x_114, 0, x_111); +lean_ctor_set(x_114, 1, x_112); +return x_114; } } } else { -uint8_t x_22; -x_22 = 0; -return x_22; -} +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_96); +lean_dec(x_23); +lean_inc(x_1); +x_115 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_90, x_1, x_21, x_97); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = lean_update_env_attributes(x_116, x_117); +if (lean_obj_tag(x_118) == 0) +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = lean_box(0); +x_122 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_119, x_121, x_120); +return x_122; } else { -uint8_t x_23; -x_23 = 0; -return x_23; +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_dec(x_2); +lean_dec(x_1); +x_123 = lean_ctor_get(x_118, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_118, 1); +lean_inc(x_124); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_125 = x_118; +} else { + lean_dec_ref(x_118); + x_125 = lean_box(0); +} +if (lean_is_scalar(x_125)) { + x_126 = lean_alloc_ctor(1, 2, 0); +} else { + x_126 = x_125; +} +lean_ctor_set(x_126, 0, x_123); +lean_ctor_set(x_126, 1, x_124); +return x_126; } } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_equivInfo___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Environment_0__Lean_equivInfo(x_1, x_2); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +lean_dec(x_93); +lean_dec(x_90); +lean_dec(x_23); +lean_dec(x_21); lean_dec(x_2); lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +x_127 = lean_ctor_get(x_95, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_95, 1); +lean_inc(x_128); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_129 = x_95; +} else { + lean_dec_ref(x_95); + x_129 = lean_box(0); } +if (lean_is_scalar(x_129)) { + x_130 = lean_alloc_ctor(1, 2, 0); +} else { + x_130 = x_129; } -LEAN_EXPORT lean_object* l_Lean_finalizeImport_unsafe__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_runtime_mark_persistent(x_1, x_2); -return x_3; +lean_ctor_set(x_130, 0, x_127); +lean_ctor_set(x_130, 1, x_128); +return x_130; } } -LEAN_EXPORT lean_object* l_Lean_finalizeImport_unsafe__2(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = lean_runtime_mark_persistent(x_1, x_2); -return x_3; +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_2); +lean_dec(x_1); +x_131 = lean_ctor_get(x_89, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_89, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_133 = x_89; +} else { + lean_dec_ref(x_89); + x_133 = lean_box(0); } +if (lean_is_scalar(x_133)) { + x_134 = lean_alloc_ctor(1, 2, 0); +} else { + x_134 = x_133; } -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; +lean_ctor_set(x_134, 0, x_131); +lean_ctor_set(x_134, 1, x_132); +return x_134; } -else -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_name_eq(x_4, x_1); -if (x_6 == 0) -{ -x_2 = x_5; -goto _start; } else { -uint8_t x_8; -x_8 = 1; -return x_8; +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_83); +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_14); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_135 = lean_ctor_get(x_84, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_84, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_137 = x_84; +} else { + lean_dec_ref(x_84); + x_137 = lean_box(0); +} +if (lean_is_scalar(x_137)) { + x_138 = lean_alloc_ctor(1, 2, 0); +} else { + x_138 = x_137; } +lean_ctor_set(x_138, 0, x_135); +lean_ctor_set(x_138, 1, x_136); +return x_138; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_finalizeImport___spec__4(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_2) == 0) +uint8_t x_139; +lean_dec(x_21); +lean_free_object(x_17); +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_139 = !lean_is_exclusive(x_22); +if (x_139 == 0) { -return x_1; +return x_22; } else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_22, 0); +x_141 = lean_ctor_get(x_22, 1); +lean_inc(x_141); +lean_inc(x_140); +lean_dec(x_22); +x_142 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +return x_142; +} +} +} +else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = l_Lean_Name_hash___override(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; -} -else +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_143 = lean_ctor_get(x_17, 0); +x_144 = lean_ctor_get(x_17, 1); +lean_inc(x_144); +lean_inc(x_143); +lean_dec(x_17); +x_145 = lean_array_get_size(x_143); +lean_dec(x_143); +x_146 = lean_get_num_attributes(x_144); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = l_Lean_Name_hash___override(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; -} +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_146, 1); +lean_inc(x_148); +lean_dec(x_146); +x_149 = lean_ctor_get(x_13, 2); +lean_inc(x_149); +lean_dec(x_13); +x_150 = lean_ctor_get(x_16, 0); +lean_inc(x_150); +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + lean_ctor_release(x_16, 1); + x_151 = x_16; +} else { + lean_dec_ref(x_16); + x_151 = lean_box(0); } +lean_inc(x_2); +lean_inc(x_4); +x_152 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_152, 0, x_4); +lean_ctor_set(x_152, 1, x_2); +lean_inc(x_150); +x_153 = lean_apply_3(x_149, x_150, x_152, x_148); +if (lean_obj_tag(x_153) == 0) +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_154 = lean_ctor_get(x_153, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_153, 1); +lean_inc(x_155); +lean_dec(x_153); +if (lean_is_scalar(x_151)) { + x_156 = lean_alloc_ctor(0, 2, 0); +} else { + x_156 = x_151; } +lean_ctor_set(x_156, 0, x_150); +lean_ctor_set(x_156, 1, x_154); +x_157 = l_Lean_EnvExtension_setState___rarg(x_14, x_4, x_156); +lean_dec(x_14); +x_158 = l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(x_157, x_155); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +lean_dec(x_158); +x_161 = lean_st_ref_get(x_6, x_160); +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = lean_ctor_get(x_161, 1); +lean_inc(x_163); +lean_dec(x_161); +x_164 = lean_get_num_attributes(x_163); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_164, 1); +lean_inc(x_166); +lean_dec(x_164); +x_167 = lean_array_get_size(x_162); +lean_dec(x_162); +x_168 = lean_nat_dec_lt(x_145, x_167); +lean_dec(x_167); +if (x_168 == 0) +{ +uint8_t x_169; +x_169 = lean_nat_dec_lt(x_147, x_165); +lean_dec(x_165); +lean_dec(x_147); +if (x_169 == 0) +{ +lean_object* x_170; lean_object* x_171; +lean_dec(x_145); +x_170 = lean_box(0); +x_171 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_159, x_170, x_166); +return x_171; } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_finalizeImport___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_inc(x_1); +x_172 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_159, x_1, x_145, x_166); +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_ctor_get(x_172, 1); +lean_inc(x_174); +lean_dec(x_172); +x_175 = lean_update_env_attributes(x_173, x_174); +if (lean_obj_tag(x_175) == 0) { -lean_dec(x_2); -lean_dec(x_1); -return x_3; +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_176 = lean_ctor_get(x_175, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_175, 1); +lean_inc(x_177); +lean_dec(x_175); +x_178 = lean_box(0); +x_179 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_176, x_178, x_177); +return x_179; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_finalizeImport___spec__4(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_dec(x_2); lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; +x_180 = lean_ctor_get(x_175, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_182 = x_175; +} else { + lean_dec_ref(x_175); + x_182 = lean_box(0); } +if (lean_is_scalar(x_182)) { + x_183 = lean_alloc_ctor(1, 2, 0); +} else { + x_183 = x_182; } +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_181); +return x_183; } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_finalizeImport___spec__3(x_7, x_1, x_6); -return x_8; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___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) { -_start: +else { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_4); -if (x_8 == 0) +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +lean_dec(x_165); +lean_dec(x_147); +lean_inc(x_1); +x_184 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_159, x_1, x_145, x_166); +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_184, 1); +lean_inc(x_186); +lean_dec(x_184); +x_187 = lean_update_env_attributes(x_185, x_186); +if (lean_obj_tag(x_187) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; size_t x_19; size_t x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; uint8_t x_25; -x_9 = lean_ctor_get(x_4, 0); -x_10 = lean_ctor_get(x_4, 1); -x_11 = lean_array_get_size(x_10); -x_12 = l_Lean_Name_hash___override(x_1); -x_13 = 32; -x_14 = lean_uint64_shift_right(x_12, x_13); -x_15 = lean_uint64_xor(x_12, x_14); -x_16 = 16; -x_17 = lean_uint64_shift_right(x_15, x_16); -x_18 = lean_uint64_xor(x_15, x_17); -x_19 = lean_uint64_to_usize(x_18); -x_20 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_21 = 1; -x_22 = lean_usize_sub(x_20, x_21); -x_23 = lean_usize_land(x_19, x_22); -x_24 = lean_array_uget(x_10, x_23); -x_25 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_1, x_24); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_26 = lean_unsigned_to_nat(1u); -x_27 = lean_nat_add(x_9, x_26); -lean_dec(x_9); -x_28 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_28, 0, x_1); -lean_ctor_set(x_28, 1, x_2); -lean_ctor_set(x_28, 2, x_24); -x_29 = lean_array_uset(x_10, x_23, x_28); -x_30 = lean_unsigned_to_nat(4u); -x_31 = lean_nat_mul(x_27, x_30); -x_32 = lean_unsigned_to_nat(3u); -x_33 = lean_nat_div(x_31, x_32); -lean_dec(x_31); -x_34 = lean_array_get_size(x_29); -x_35 = lean_nat_dec_le(x_33, x_34); -lean_dec(x_34); -lean_dec(x_33); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(x_29); -lean_ctor_set(x_4, 1, x_36); -lean_ctor_set(x_4, 0, x_27); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_4); -lean_ctor_set(x_37, 1, x_5); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_3); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_7); -return x_40; +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_188 = lean_ctor_get(x_187, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_187, 1); +lean_inc(x_189); +lean_dec(x_187); +x_190 = lean_box(0); +x_191 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_188, x_190, x_189); +return x_191; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_ctor_set(x_4, 1, x_29); -lean_ctor_set(x_4, 0, x_27); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_4); -lean_ctor_set(x_41, 1, x_5); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_3); -lean_ctor_set(x_42, 1, x_41); -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_42); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_7); -return x_44; +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +lean_dec(x_2); +lean_dec(x_1); +x_192 = lean_ctor_get(x_187, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_187, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_194 = x_187; +} else { + lean_dec_ref(x_187); + x_194 = lean_box(0); +} +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); +} else { + x_195 = x_194; +} +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_193); +return x_195; +} } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_24); +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +lean_dec(x_162); +lean_dec(x_159); +lean_dec(x_147); +lean_dec(x_145); lean_dec(x_2); lean_dec(x_1); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_4); -lean_ctor_set(x_45, 1, x_5); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_3); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_46); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_7); -return x_48; -} +x_196 = lean_ctor_get(x_164, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_164, 1); +lean_inc(x_197); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_198 = x_164; +} else { + lean_dec_ref(x_164); + x_198 = lean_box(0); } -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; size_t x_59; size_t x_60; size_t x_61; size_t x_62; size_t x_63; lean_object* x_64; uint8_t x_65; -x_49 = lean_ctor_get(x_4, 0); -x_50 = lean_ctor_get(x_4, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_4); -x_51 = lean_array_get_size(x_50); -x_52 = l_Lean_Name_hash___override(x_1); -x_53 = 32; -x_54 = lean_uint64_shift_right(x_52, x_53); -x_55 = lean_uint64_xor(x_52, x_54); -x_56 = 16; -x_57 = lean_uint64_shift_right(x_55, x_56); -x_58 = lean_uint64_xor(x_55, x_57); -x_59 = lean_uint64_to_usize(x_58); -x_60 = lean_usize_of_nat(x_51); -lean_dec(x_51); -x_61 = 1; -x_62 = lean_usize_sub(x_60, x_61); -x_63 = lean_usize_land(x_59, x_62); -x_64 = lean_array_uget(x_50, x_63); -x_65 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_1, 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; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; -x_66 = lean_unsigned_to_nat(1u); -x_67 = lean_nat_add(x_49, x_66); -lean_dec(x_49); -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_2); -lean_ctor_set(x_68, 2, x_64); -x_69 = lean_array_uset(x_50, x_63, x_68); -x_70 = lean_unsigned_to_nat(4u); -x_71 = lean_nat_mul(x_67, x_70); -x_72 = lean_unsigned_to_nat(3u); -x_73 = lean_nat_div(x_71, x_72); -lean_dec(x_71); -x_74 = lean_array_get_size(x_69); -x_75 = lean_nat_dec_le(x_73, x_74); -lean_dec(x_74); -lean_dec(x_73); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_76 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(x_69); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_67); -lean_ctor_set(x_77, 1, x_76); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_5); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_3); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_80, 0, x_79); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_7); -return x_81; +if (lean_is_scalar(x_198)) { + x_199 = lean_alloc_ctor(1, 2, 0); +} else { + x_199 = x_198; } -else -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_67); -lean_ctor_set(x_82, 1, x_69); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_5); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_3); -lean_ctor_set(x_84, 1, x_83); -x_85 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_85, 0, 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_7); -return x_86; +lean_ctor_set(x_199, 0, x_196); +lean_ctor_set(x_199, 1, x_197); +return x_199; } } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -lean_dec(x_64); +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +lean_dec(x_147); +lean_dec(x_145); lean_dec(x_2); lean_dec(x_1); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_49); -lean_ctor_set(x_87, 1, x_50); -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_5); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_3); -lean_ctor_set(x_89, 1, x_88); -x_90 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_90, 0, x_89); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_7); -return x_91; +x_200 = lean_ctor_get(x_158, 0); +lean_inc(x_200); +x_201 = lean_ctor_get(x_158, 1); +lean_inc(x_201); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_202 = x_158; +} else { + lean_dec_ref(x_158); + x_202 = lean_box(0); } +if (lean_is_scalar(x_202)) { + x_203 = lean_alloc_ctor(1, 2, 0); +} else { + x_203 = x_202; } +lean_ctor_set(x_203, 0, x_200); +lean_ctor_set(x_203, 1, x_201); +return x_203; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_lt(x_7, x_6); -if (x_10 == 0) +else { -lean_object* x_11; -lean_dec(x_2); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_8); -lean_ctor_set(x_11, 1, x_9); -return x_11; +lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +lean_dec(x_151); +lean_dec(x_150); +lean_dec(x_147); +lean_dec(x_145); +lean_dec(x_14); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_204 = lean_ctor_get(x_153, 0); +lean_inc(x_204); +x_205 = lean_ctor_get(x_153, 1); +lean_inc(x_205); +if (lean_is_exclusive(x_153)) { + lean_ctor_release(x_153, 0); + lean_ctor_release(x_153, 1); + x_206 = x_153; +} else { + lean_dec_ref(x_153); + x_206 = lean_box(0); +} +if (lean_is_scalar(x_206)) { + x_207 = lean_alloc_ctor(1, 2, 0); +} else { + x_207 = x_206; +} +lean_ctor_set(x_207, 0, x_204); +lean_ctor_set(x_207, 1, x_205); +return x_207; +} } else { -lean_object* x_12; uint8_t x_13; -x_12 = lean_array_uget(x_5, x_7); -x_13 = !lean_is_exclusive(x_8); -if (x_13 == 0) -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_8, 1); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_16 = lean_ctor_get(x_8, 0); -x_17 = lean_ctor_get(x_14, 0); -x_18 = lean_ctor_get(x_14, 1); -x_19 = lean_ctor_get(x_16, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_16, 1); -lean_inc(x_20); -x_21 = lean_ctor_get(x_16, 2); -lean_inc(x_21); -x_22 = lean_nat_dec_lt(x_20, x_21); -if (x_22 == 0) -{ -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_12); +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +lean_dec(x_145); +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_4); lean_dec(x_2); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_8); -lean_ctor_set(x_23, 1, x_9); -return x_23; +lean_dec(x_1); +x_208 = lean_ctor_get(x_146, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_146, 1); +lean_inc(x_209); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_210 = x_146; +} else { + lean_dec_ref(x_146); + x_210 = lean_box(0); +} +if (lean_is_scalar(x_210)) { + x_211 = lean_alloc_ctor(1, 2, 0); +} else { + x_211 = x_210; +} +lean_ctor_set(x_211, 0, x_208); +lean_ctor_set(x_211, 1, x_209); +return x_211; } -else -{ -uint8_t x_24; -lean_free_object(x_14); -lean_free_object(x_8); -x_24 = !lean_is_exclusive(x_16); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_57; -x_25 = lean_ctor_get(x_16, 2); -lean_dec(x_25); -x_26 = lean_ctor_get(x_16, 1); -lean_dec(x_26); -x_27 = lean_ctor_get(x_16, 0); -lean_dec(x_27); -x_28 = lean_array_fget(x_19, x_20); -x_29 = lean_unsigned_to_nat(1u); -x_30 = lean_nat_add(x_20, x_29); -lean_dec(x_20); -lean_ctor_set(x_16, 1, x_30); -x_57 = !lean_is_exclusive(x_18); -if (x_57 == 0) -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; uint64_t x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; size_t x_68; size_t x_69; size_t x_70; size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; -x_58 = lean_ctor_get(x_18, 0); -x_59 = lean_ctor_get(x_18, 1); -x_60 = lean_array_get_size(x_59); -x_61 = l_Lean_Name_hash___override(x_12); -x_62 = 32; -x_63 = lean_uint64_shift_right(x_61, x_62); -x_64 = lean_uint64_xor(x_61, x_63); -x_65 = 16; -x_66 = lean_uint64_shift_right(x_64, x_65); -x_67 = lean_uint64_xor(x_64, x_66); -x_68 = lean_uint64_to_usize(x_67); -x_69 = lean_usize_of_nat(x_60); -lean_dec(x_60); -x_70 = 1; -x_71 = lean_usize_sub(x_69, x_70); -x_72 = lean_usize_land(x_68, x_71); -x_73 = lean_array_uget(x_59, x_72); -x_74 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_12, x_73); -if (lean_obj_tag(x_74) == 0) -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; -x_75 = lean_nat_add(x_58, x_29); -lean_dec(x_58); -lean_inc(x_28); -lean_inc(x_12); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_12); -lean_ctor_set(x_76, 1, x_28); -lean_ctor_set(x_76, 2, x_73); -x_77 = lean_array_uset(x_59, x_72, x_76); -x_78 = lean_box(0); -x_79 = lean_unsigned_to_nat(4u); -x_80 = lean_nat_mul(x_75, x_79); -x_81 = lean_unsigned_to_nat(3u); -x_82 = lean_nat_div(x_80, x_81); -lean_dec(x_80); -x_83 = lean_array_get_size(x_77); -x_84 = lean_nat_dec_le(x_82, x_83); -lean_dec(x_83); -lean_dec(x_82); -if (x_84 == 0) -{ -lean_object* x_85; -x_85 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_77); -lean_ctor_set(x_18, 1, x_85); -lean_ctor_set(x_18, 0, x_75); -x_31 = x_78; -x_32 = x_18; -goto block_56; } -else -{ -lean_ctor_set(x_18, 1, x_77); -lean_ctor_set(x_18, 0, x_75); -x_31 = x_78; -x_32 = x_18; -goto block_56; } } else { -uint8_t x_86; -lean_dec(x_73); -x_86 = !lean_is_exclusive(x_74); -if (x_86 == 0) +lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; +x_212 = lean_ctor_get(x_7, 0); +x_213 = lean_ctor_get(x_7, 1); +lean_inc(x_213); +lean_inc(x_212); +lean_dec(x_7); +x_214 = lean_array_get_size(x_212); +x_215 = lean_nat_dec_lt(x_3, x_214); +lean_dec(x_214); +if (x_215 == 0) { -x_31 = x_74; -x_32 = x_18; -goto block_56; +lean_object* x_216; +lean_dec(x_212); +lean_dec(x_2); +lean_dec(x_1); +x_216 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_216, 0, x_4); +lean_ctor_set(x_216, 1, x_213); +return x_216; } else { -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_74, 0); -lean_inc(x_87); -lean_dec(x_74); -x_88 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_88, 0, x_87); -x_31 = x_88; -x_32 = x_18; -goto block_56; +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_217 = lean_array_fget(x_212, x_3); +lean_dec(x_212); +x_218 = lean_ctor_get(x_217, 0); +lean_inc(x_218); +x_219 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1; +x_220 = l_Lean_EnvExtension_getState___rarg(x_219, x_218, x_4); +x_221 = lean_st_ref_get(x_6, x_213); +x_222 = lean_ctor_get(x_221, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_221, 1); +lean_inc(x_223); +if (lean_is_exclusive(x_221)) { + lean_ctor_release(x_221, 0); + lean_ctor_release(x_221, 1); + x_224 = x_221; +} else { + lean_dec_ref(x_221); + x_224 = lean_box(0); } +x_225 = lean_array_get_size(x_222); +lean_dec(x_222); +x_226 = lean_get_num_attributes(x_223); +if (lean_obj_tag(x_226) == 0) +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_226, 1); +lean_inc(x_228); +lean_dec(x_226); +x_229 = lean_ctor_get(x_217, 2); +lean_inc(x_229); +lean_dec(x_217); +x_230 = lean_ctor_get(x_220, 0); +lean_inc(x_230); +if (lean_is_exclusive(x_220)) { + lean_ctor_release(x_220, 0); + lean_ctor_release(x_220, 1); + x_231 = x_220; +} else { + lean_dec_ref(x_220); + x_231 = lean_box(0); } +lean_inc(x_2); +lean_inc(x_4); +if (lean_is_scalar(x_224)) { + x_232 = lean_alloc_ctor(0, 2, 0); +} else { + x_232 = x_224; } -else +lean_ctor_set(x_232, 0, x_4); +lean_ctor_set(x_232, 1, x_2); +lean_inc(x_230); +x_233 = lean_apply_3(x_229, x_230, x_232, x_228); +if (lean_obj_tag(x_233) == 0) { -lean_object* x_89; lean_object* x_90; lean_object* x_91; uint64_t x_92; uint64_t x_93; uint64_t x_94; uint64_t x_95; uint64_t x_96; uint64_t x_97; uint64_t x_98; size_t x_99; size_t x_100; size_t x_101; size_t x_102; size_t x_103; lean_object* x_104; lean_object* x_105; -x_89 = lean_ctor_get(x_18, 0); -x_90 = lean_ctor_get(x_18, 1); -lean_inc(x_90); -lean_inc(x_89); -lean_dec(x_18); -x_91 = lean_array_get_size(x_90); -x_92 = l_Lean_Name_hash___override(x_12); -x_93 = 32; -x_94 = lean_uint64_shift_right(x_92, x_93); -x_95 = lean_uint64_xor(x_92, x_94); -x_96 = 16; -x_97 = lean_uint64_shift_right(x_95, x_96); -x_98 = lean_uint64_xor(x_95, x_97); -x_99 = lean_uint64_to_usize(x_98); -x_100 = lean_usize_of_nat(x_91); -lean_dec(x_91); -x_101 = 1; -x_102 = lean_usize_sub(x_100, x_101); -x_103 = lean_usize_land(x_99, x_102); -x_104 = lean_array_uget(x_90, x_103); -x_105 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_12, x_104); -if (lean_obj_tag(x_105) == 0) +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_233, 1); +lean_inc(x_235); +lean_dec(x_233); +if (lean_is_scalar(x_231)) { + x_236 = lean_alloc_ctor(0, 2, 0); +} else { + x_236 = x_231; +} +lean_ctor_set(x_236, 0, x_230); +lean_ctor_set(x_236, 1, x_234); +x_237 = l_Lean_EnvExtension_setState___rarg(x_218, x_4, x_236); +lean_dec(x_218); +x_238 = l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(x_237, x_235); +if (lean_obj_tag(x_238) == 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; uint8_t x_115; -x_106 = lean_nat_add(x_89, x_29); -lean_dec(x_89); -lean_inc(x_28); -lean_inc(x_12); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_12); -lean_ctor_set(x_107, 1, x_28); -lean_ctor_set(x_107, 2, x_104); -x_108 = lean_array_uset(x_90, x_103, x_107); -x_109 = lean_box(0); -x_110 = lean_unsigned_to_nat(4u); -x_111 = lean_nat_mul(x_106, x_110); -x_112 = lean_unsigned_to_nat(3u); -x_113 = lean_nat_div(x_111, x_112); -lean_dec(x_111); -x_114 = lean_array_get_size(x_108); -x_115 = lean_nat_dec_le(x_113, x_114); -lean_dec(x_114); -lean_dec(x_113); -if (x_115 == 0) +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; +x_239 = lean_ctor_get(x_238, 0); +lean_inc(x_239); +x_240 = lean_ctor_get(x_238, 1); +lean_inc(x_240); +lean_dec(x_238); +x_241 = lean_st_ref_get(x_6, x_240); +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); +x_244 = lean_get_num_attributes(x_243); +if (lean_obj_tag(x_244) == 0) { -lean_object* x_116; lean_object* x_117; -x_116 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_108); -x_117 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_117, 0, x_106); -lean_ctor_set(x_117, 1, x_116); -x_31 = x_109; -x_32 = x_117; -goto block_56; +lean_object* x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; +x_245 = lean_ctor_get(x_244, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_244, 1); +lean_inc(x_246); +lean_dec(x_244); +x_247 = lean_array_get_size(x_242); +lean_dec(x_242); +x_248 = lean_nat_dec_lt(x_225, x_247); +lean_dec(x_247); +if (x_248 == 0) +{ +uint8_t x_249; +x_249 = lean_nat_dec_lt(x_227, x_245); +lean_dec(x_245); +lean_dec(x_227); +if (x_249 == 0) +{ +lean_object* x_250; lean_object* x_251; +lean_dec(x_225); +x_250 = lean_box(0); +x_251 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_239, x_250, x_246); +return x_251; } else { -lean_object* x_118; -x_118 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_118, 0, x_106); -lean_ctor_set(x_118, 1, x_108); -x_31 = x_109; -x_32 = x_118; -goto block_56; -} +lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; +lean_inc(x_1); +x_252 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_239, x_1, x_225, x_246); +x_253 = lean_ctor_get(x_252, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_252, 1); +lean_inc(x_254); +lean_dec(x_252); +x_255 = lean_update_env_attributes(x_253, x_254); +if (lean_obj_tag(x_255) == 0) +{ +lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; +x_256 = lean_ctor_get(x_255, 0); +lean_inc(x_256); +x_257 = lean_ctor_get(x_255, 1); +lean_inc(x_257); +lean_dec(x_255); +x_258 = lean_box(0); +x_259 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_256, x_258, x_257); +return x_259; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_dec(x_104); -x_119 = lean_ctor_get(x_105, 0); -lean_inc(x_119); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - x_120 = x_105; +lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +lean_dec(x_2); +lean_dec(x_1); +x_260 = lean_ctor_get(x_255, 0); +lean_inc(x_260); +x_261 = lean_ctor_get(x_255, 1); +lean_inc(x_261); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_262 = x_255; } else { - lean_dec_ref(x_105); - x_120 = lean_box(0); + lean_dec_ref(x_255); + x_262 = lean_box(0); } -if (lean_is_scalar(x_120)) { - x_121 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_262)) { + x_263 = lean_alloc_ctor(1, 2, 0); } else { - x_121 = x_120; + x_263 = x_262; } -lean_ctor_set(x_121, 0, x_119); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_89); -lean_ctor_set(x_122, 1, x_90); -x_31 = x_121; -x_32 = x_122; -goto block_56; +lean_ctor_set(x_263, 0, x_260); +lean_ctor_set(x_263, 1, x_261); +return x_263; } } -block_56: +} +else { -if (lean_obj_tag(x_31) == 0) +lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; +lean_dec(x_245); +lean_dec(x_227); +lean_inc(x_1); +x_264 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_239, x_1, x_225, x_246); +x_265 = lean_ctor_get(x_264, 0); +lean_inc(x_265); +x_266 = lean_ctor_get(x_264, 1); +lean_inc(x_266); +lean_dec(x_264); +x_267 = lean_update_env_attributes(x_265, x_266); +if (lean_obj_tag(x_267) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; size_t x_39; -lean_dec(x_28); -x_33 = lean_box(0); -lean_inc(x_2); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_16, x_17, x_32, x_33, x_9); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_ctor_get(x_35, 0); -lean_inc(x_37); -lean_dec(x_35); -x_38 = 1; -x_39 = lean_usize_add(x_7, x_38); -x_7 = x_39; -x_8 = x_37; -x_9 = x_36; -goto _start; +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +x_268 = lean_ctor_get(x_267, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_267, 1); +lean_inc(x_269); +lean_dec(x_267); +x_270 = lean_box(0); +x_271 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_3, x_1, x_2, x_268, x_270, x_269); +return x_271; } else { -lean_object* x_41; uint8_t x_42; -x_41 = lean_ctor_get(x_31, 0); -lean_inc(x_41); -lean_dec(x_31); -x_42 = l___private_Lean_Environment_0__Lean_equivInfo(x_41, x_28); -lean_dec(x_28); -lean_dec(x_41); -if (x_42 == 0) -{ -lean_object* x_43; uint8_t x_44; -lean_dec(x_32); -lean_dec(x_16); -x_43 = l_Lean_throwAlreadyImported___rarg(x_1, x_17, x_2, x_12, x_9); +lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_dec(x_2); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) -{ -return x_43; +lean_dec(x_1); +x_272 = lean_ctor_get(x_267, 0); +lean_inc(x_272); +x_273 = lean_ctor_get(x_267, 1); +lean_inc(x_273); +if (lean_is_exclusive(x_267)) { + lean_ctor_release(x_267, 0); + lean_ctor_release(x_267, 1); + x_274 = x_267; +} else { + lean_dec_ref(x_267); + x_274 = lean_box(0); +} +if (lean_is_scalar(x_274)) { + x_275 = lean_alloc_ctor(1, 2, 0); +} else { + x_275 = x_274; +} +lean_ctor_set(x_275, 0, x_272); +lean_ctor_set(x_275, 1, x_273); +return x_275; +} +} } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_43, 0); -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_43); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; -} -} -else -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; -x_48 = lean_box(0); -lean_inc(x_2); -x_49 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_16, x_17, x_32, x_48, x_9); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_ctor_get(x_50, 0); -lean_inc(x_52); -lean_dec(x_50); -x_53 = 1; -x_54 = lean_usize_add(x_7, x_53); -x_7 = x_54; -x_8 = x_52; -x_9 = x_51; -goto _start; +lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; +lean_dec(x_242); +lean_dec(x_239); +lean_dec(x_227); +lean_dec(x_225); +lean_dec(x_2); +lean_dec(x_1); +x_276 = lean_ctor_get(x_244, 0); +lean_inc(x_276); +x_277 = lean_ctor_get(x_244, 1); +lean_inc(x_277); +if (lean_is_exclusive(x_244)) { + lean_ctor_release(x_244, 0); + lean_ctor_release(x_244, 1); + x_278 = x_244; +} else { + lean_dec_ref(x_244); + x_278 = lean_box(0); } +if (lean_is_scalar(x_278)) { + x_279 = lean_alloc_ctor(1, 2, 0); +} else { + x_279 = x_278; } +lean_ctor_set(x_279, 0, x_276); +lean_ctor_set(x_279, 1, x_277); +return x_279; } } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint64_t x_157; uint64_t x_158; uint64_t x_159; uint64_t x_160; uint64_t x_161; uint64_t x_162; uint64_t x_163; size_t x_164; size_t x_165; size_t x_166; size_t x_167; size_t x_168; lean_object* x_169; lean_object* x_170; -lean_dec(x_16); -x_123 = lean_array_fget(x_19, x_20); -x_124 = lean_unsigned_to_nat(1u); -x_125 = lean_nat_add(x_20, x_124); -lean_dec(x_20); -x_126 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_126, 0, x_19); -lean_ctor_set(x_126, 1, x_125); -lean_ctor_set(x_126, 2, x_21); -x_153 = lean_ctor_get(x_18, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_18, 1); -lean_inc(x_154); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - lean_ctor_release(x_18, 1); - x_155 = x_18; +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +lean_dec(x_227); +lean_dec(x_225); +lean_dec(x_2); +lean_dec(x_1); +x_280 = lean_ctor_get(x_238, 0); +lean_inc(x_280); +x_281 = lean_ctor_get(x_238, 1); +lean_inc(x_281); +if (lean_is_exclusive(x_238)) { + lean_ctor_release(x_238, 0); + lean_ctor_release(x_238, 1); + x_282 = x_238; } else { - lean_dec_ref(x_18); - x_155 = lean_box(0); + lean_dec_ref(x_238); + x_282 = lean_box(0); } -x_156 = lean_array_get_size(x_154); -x_157 = l_Lean_Name_hash___override(x_12); -x_158 = 32; -x_159 = lean_uint64_shift_right(x_157, x_158); -x_160 = lean_uint64_xor(x_157, x_159); -x_161 = 16; -x_162 = lean_uint64_shift_right(x_160, x_161); -x_163 = lean_uint64_xor(x_160, x_162); -x_164 = lean_uint64_to_usize(x_163); -x_165 = lean_usize_of_nat(x_156); -lean_dec(x_156); -x_166 = 1; -x_167 = lean_usize_sub(x_165, x_166); -x_168 = lean_usize_land(x_164, x_167); -x_169 = lean_array_uget(x_154, x_168); -x_170 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_12, x_169); -if (lean_obj_tag(x_170) == 0) -{ -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; -x_171 = lean_nat_add(x_153, x_124); -lean_dec(x_153); -lean_inc(x_123); -lean_inc(x_12); -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_12); -lean_ctor_set(x_172, 1, x_123); -lean_ctor_set(x_172, 2, x_169); -x_173 = lean_array_uset(x_154, x_168, x_172); -x_174 = lean_box(0); -x_175 = lean_unsigned_to_nat(4u); -x_176 = lean_nat_mul(x_171, x_175); -x_177 = lean_unsigned_to_nat(3u); -x_178 = lean_nat_div(x_176, x_177); -lean_dec(x_176); -x_179 = lean_array_get_size(x_173); -x_180 = lean_nat_dec_le(x_178, x_179); -lean_dec(x_179); -lean_dec(x_178); -if (x_180 == 0) -{ -lean_object* x_181; lean_object* x_182; -x_181 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_173); -if (lean_is_scalar(x_155)) { - x_182 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_282)) { + x_283 = lean_alloc_ctor(1, 2, 0); } else { - x_182 = x_155; + x_283 = x_282; +} +lean_ctor_set(x_283, 0, x_280); +lean_ctor_set(x_283, 1, x_281); +return x_283; } -lean_ctor_set(x_182, 0, x_171); -lean_ctor_set(x_182, 1, x_181); -x_127 = x_174; -x_128 = x_182; -goto block_152; } else { -lean_object* x_183; -if (lean_is_scalar(x_155)) { - x_183 = lean_alloc_ctor(0, 2, 0); +lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; +lean_dec(x_231); +lean_dec(x_230); +lean_dec(x_227); +lean_dec(x_225); +lean_dec(x_218); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_284 = lean_ctor_get(x_233, 0); +lean_inc(x_284); +x_285 = lean_ctor_get(x_233, 1); +lean_inc(x_285); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_286 = x_233; } else { - x_183 = x_155; + lean_dec_ref(x_233); + x_286 = lean_box(0); } -lean_ctor_set(x_183, 0, x_171); -lean_ctor_set(x_183, 1, x_173); -x_127 = x_174; -x_128 = x_183; -goto block_152; +if (lean_is_scalar(x_286)) { + x_287 = lean_alloc_ctor(1, 2, 0); +} else { + x_287 = x_286; +} +lean_ctor_set(x_287, 0, x_284); +lean_ctor_set(x_287, 1, x_285); +return x_287; } } else { -lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -lean_dec(x_169); -x_184 = lean_ctor_get(x_170, 0); -lean_inc(x_184); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - x_185 = x_170; +lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; +lean_dec(x_225); +lean_dec(x_224); +lean_dec(x_220); +lean_dec(x_218); +lean_dec(x_217); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_288 = lean_ctor_get(x_226, 0); +lean_inc(x_288); +x_289 = lean_ctor_get(x_226, 1); +lean_inc(x_289); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_290 = x_226; } else { - lean_dec_ref(x_170); - x_185 = lean_box(0); + lean_dec_ref(x_226); + x_290 = lean_box(0); } -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_290)) { + x_291 = lean_alloc_ctor(1, 2, 0); } else { - x_186 = x_185; + x_291 = x_290; } -lean_ctor_set(x_186, 0, x_184); -if (lean_is_scalar(x_155)) { - x_187 = lean_alloc_ctor(0, 2, 0); -} else { - x_187 = x_155; +lean_ctor_set(x_291, 0, x_288); +lean_ctor_set(x_291, 1, x_289); +return x_291; } -lean_ctor_set(x_187, 0, x_153); -lean_ctor_set(x_187, 1, x_154); -x_127 = x_186; -x_128 = x_187; -goto block_152; } -block_152: +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -if (lean_obj_tag(x_127) == 0) +lean_object* x_7; +x_7 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___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_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; size_t x_134; size_t x_135; -lean_dec(x_123); -x_129 = lean_box(0); -lean_inc(x_2); -x_130 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_126, x_17, x_128, x_129, x_9); -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_130, 1); -lean_inc(x_132); -lean_dec(x_130); -x_133 = lean_ctor_get(x_131, 0); -lean_inc(x_133); -lean_dec(x_131); -x_134 = 1; -x_135 = lean_usize_add(x_7, x_134); -x_7 = x_135; -x_8 = x_133; -x_9 = x_132; -goto _start; +lean_object* x_6; +x_6 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +return x_6; } -else -{ -lean_object* x_137; uint8_t x_138; -x_137 = lean_ctor_get(x_127, 0); -lean_inc(x_137); -lean_dec(x_127); -x_138 = l___private_Lean_Environment_0__Lean_equivInfo(x_137, x_123); -lean_dec(x_123); -lean_dec(x_137); -if (x_138 == 0) +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -lean_dec(x_128); -lean_dec(x_126); -x_139 = l_Lean_throwAlreadyImported___rarg(x_1, x_17, x_2, x_12, x_9); -lean_dec(x_2); -x_140 = lean_ctor_get(x_139, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_139, 1); -lean_inc(x_141); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_142 = x_139; -} else { - lean_dec_ref(x_139); - x_142 = lean_box(0); +lean_object* x_5; lean_object* x_6; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(x_2, x_3, x_5, x_1, x_4); +return x_6; } -if (lean_is_scalar(x_142)) { - x_143 = lean_alloc_ctor(1, 2, 0); -} else { - x_143 = x_142; } -lean_ctor_set(x_143, 0, x_140); -lean_ctor_set(x_143, 1, x_141); -return x_143; +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Data.DHashMap.Internal.AssocList.Basic", 42, 42); +return x_1; } -else +} +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__2() { +_start: { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; size_t x_149; size_t x_150; -x_144 = lean_box(0); -lean_inc(x_2); -x_145 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_126, x_17, x_128, x_144, x_9); -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -lean_dec(x_145); -x_148 = lean_ctor_get(x_146, 0); -lean_inc(x_148); -lean_dec(x_146); -x_149 = 1; -x_150 = lean_usize_add(x_7, x_149); -x_7 = x_150; -x_8 = x_148; -x_9 = x_147; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.DHashMap.Internal.AssocList.get!", 36, 36); +return x_1; +} } +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("key is not present in hash table", 32, 32); +return x_1; } } +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__1; +x_2 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__2; +x_3 = lean_unsigned_to_nat(138u); +x_4 = lean_unsigned_to_nat(11u); +x_5 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___closed__4; +x_5 = l_panic___rarg(x_1, x_4); +return x_5; } else { -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; uint8_t x_194; -x_188 = lean_ctor_get(x_8, 0); -x_189 = lean_ctor_get(x_14, 0); -x_190 = lean_ctor_get(x_14, 1); -lean_inc(x_190); -lean_inc(x_189); -lean_dec(x_14); -x_191 = lean_ctor_get(x_188, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_188, 1); -lean_inc(x_192); -x_193 = lean_ctor_get(x_188, 2); -lean_inc(x_193); -x_194 = lean_nat_dec_lt(x_192, x_193); -if (x_194 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_name_eq(x_6, x_2); +if (x_9 == 0) { -lean_object* x_195; lean_object* x_196; -lean_dec(x_193); -lean_dec(x_192); -lean_dec(x_191); -lean_dec(x_12); -lean_dec(x_2); -x_195 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_195, 0, x_189); -lean_ctor_set(x_195, 1, x_190); -lean_ctor_set(x_8, 1, x_195); -x_196 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_196, 0, x_8); -lean_ctor_set(x_196, 1, x_9); -return x_196; +x_3 = x_8; +goto _start; } else { -lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; uint64_t x_232; uint64_t x_233; uint64_t x_234; uint64_t x_235; uint64_t x_236; uint64_t x_237; uint64_t x_238; size_t x_239; size_t x_240; size_t x_241; size_t x_242; size_t x_243; lean_object* x_244; lean_object* x_245; -lean_free_object(x_8); -if (lean_is_exclusive(x_188)) { - lean_ctor_release(x_188, 0); - lean_ctor_release(x_188, 1); - lean_ctor_release(x_188, 2); - x_197 = x_188; -} else { - lean_dec_ref(x_188); - x_197 = lean_box(0); +lean_dec(x_1); +lean_inc(x_7); +return x_7; } -x_198 = lean_array_fget(x_191, x_192); -x_199 = lean_unsigned_to_nat(1u); -x_200 = lean_nat_add(x_192, x_199); -lean_dec(x_192); -if (lean_is_scalar(x_197)) { - x_201 = lean_alloc_ctor(0, 3, 0); -} else { - x_201 = x_197; } -lean_ctor_set(x_201, 0, x_191); -lean_ctor_set(x_201, 1, x_200); -lean_ctor_set(x_201, 2, x_193); -x_228 = lean_ctor_get(x_190, 0); -lean_inc(x_228); -x_229 = lean_ctor_get(x_190, 1); -lean_inc(x_229); -if (lean_is_exclusive(x_190)) { - lean_ctor_release(x_190, 0); - lean_ctor_release(x_190, 1); - x_230 = x_190; -} else { - lean_dec_ref(x_190); - x_230 = lean_box(0); } -x_231 = lean_array_get_size(x_229); -x_232 = l_Lean_Name_hash___override(x_12); -x_233 = 32; -x_234 = lean_uint64_shift_right(x_232, x_233); -x_235 = lean_uint64_xor(x_232, x_234); -x_236 = 16; -x_237 = lean_uint64_shift_right(x_235, x_236); -x_238 = lean_uint64_xor(x_235, x_237); -x_239 = lean_uint64_to_usize(x_238); -x_240 = lean_usize_of_nat(x_231); -lean_dec(x_231); -x_241 = 1; -x_242 = lean_usize_sub(x_240, x_241); -x_243 = lean_usize_land(x_239, x_242); -x_244 = lean_array_uget(x_229, x_243); -x_245 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_12, x_244); -if (lean_obj_tag(x_245) == 0) -{ -lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; uint8_t x_255; -x_246 = lean_nat_add(x_228, x_199); -lean_dec(x_228); -lean_inc(x_198); -lean_inc(x_12); -x_247 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_247, 0, x_12); -lean_ctor_set(x_247, 1, x_198); -lean_ctor_set(x_247, 2, x_244); -x_248 = lean_array_uset(x_229, x_243, x_247); -x_249 = lean_box(0); -x_250 = lean_unsigned_to_nat(4u); -x_251 = lean_nat_mul(x_246, x_250); -x_252 = lean_unsigned_to_nat(3u); -x_253 = lean_nat_div(x_251, x_252); -lean_dec(x_251); -x_254 = lean_array_get_size(x_248); -x_255 = lean_nat_dec_le(x_253, x_254); -lean_dec(x_254); -lean_dec(x_253); -if (x_255 == 0) +} +static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__1() { +_start: { -lean_object* x_256; lean_object* x_257; -x_256 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_248); -if (lean_is_scalar(x_230)) { - x_257 = lean_alloc_ctor(0, 2, 0); -} else { - x_257 = x_230; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("import ", 7, 7); +return x_1; } -lean_ctor_set(x_257, 0, x_246); -lean_ctor_set(x_257, 1, x_256); -x_202 = x_249; -x_203 = x_257; -goto block_227; } -else +static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__2() { +_start: { -lean_object* x_258; -if (lean_is_scalar(x_230)) { - x_258 = lean_alloc_ctor(0, 2, 0); -} else { - x_258 = x_230; -} -lean_ctor_set(x_258, 0, x_246); -lean_ctor_set(x_258, 1, x_248); -x_202 = x_249; -x_203 = x_258; -goto block_227; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" failed, environment already contains '", 39, 39); +return x_1; } } -else +static lean_object* _init_l_Lean_throwAlreadyImported___rarg___closed__3() { +_start: { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -lean_dec(x_244); -x_259 = lean_ctor_get(x_245, 0); -lean_inc(x_259); -if (lean_is_exclusive(x_245)) { - lean_ctor_release(x_245, 0); - x_260 = x_245; -} else { - lean_dec_ref(x_245); - x_260 = lean_box(0); -} -if (lean_is_scalar(x_260)) { - x_261 = lean_alloc_ctor(1, 1, 0); -} else { - x_261 = x_260; -} -lean_ctor_set(x_261, 0, x_259); -if (lean_is_scalar(x_230)) { - x_262 = lean_alloc_ctor(0, 2, 0); -} else { - x_262 = x_230; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' from ", 7, 7); +return x_1; } -lean_ctor_set(x_262, 0, x_228); -lean_ctor_set(x_262, 1, x_229); -x_202 = x_261; -x_203 = x_262; -goto block_227; } -block_227: +LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -if (lean_obj_tag(x_202) == 0) +lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_array_get_size(x_6); +x_8 = lean_nat_dec_lt(x_3, x_7); +x_9 = 1; +x_10 = l_Lean_instToStringImport___closed__1; +lean_inc(x_4); +x_11 = l_Lean_Name_toString(x_4, x_9, x_10); +if (x_8 == 0) { -lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; size_t x_209; size_t x_210; -lean_dec(x_198); -x_204 = lean_box(0); -lean_inc(x_2); -x_205 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_201, x_189, x_203, x_204, x_9); -x_206 = lean_ctor_get(x_205, 0); -lean_inc(x_206); -x_207 = lean_ctor_get(x_205, 1); -lean_inc(x_207); -lean_dec(x_205); -x_208 = lean_ctor_get(x_206, 0); -lean_inc(x_208); -lean_dec(x_206); -x_209 = 1; -x_210 = lean_usize_add(x_7, x_209); -x_7 = x_210; -x_8 = x_208; -x_9 = x_207; -goto _start; +lean_object* x_88; lean_object* x_89; +x_88 = l_Lean_instInhabitedName; +x_89 = l_outOfBounds___rarg(x_88); +x_12 = x_89; +goto block_87; } else { -lean_object* x_212; uint8_t x_213; -x_212 = lean_ctor_get(x_202, 0); -lean_inc(x_212); -lean_dec(x_202); -x_213 = l___private_Lean_Environment_0__Lean_equivInfo(x_212, x_198); -lean_dec(x_198); -lean_dec(x_212); -if (x_213 == 0) +lean_object* x_90; +x_90 = lean_array_fget(x_6, x_3); +x_12 = x_90; +goto block_87; +} +block_87: { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; -lean_dec(x_203); -lean_dec(x_201); -x_214 = l_Lean_throwAlreadyImported___rarg(x_1, x_189, x_2, x_12, x_9); -lean_dec(x_2); -x_215 = lean_ctor_get(x_214, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_214, 1); -lean_inc(x_216); -if (lean_is_exclusive(x_214)) { - lean_ctor_release(x_214, 0); - lean_ctor_release(x_214, 1); - x_217 = x_214; -} else { - lean_dec_ref(x_214); - x_217 = lean_box(0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_13 = l_Lean_Name_toString(x_12, x_9, x_10); +x_14 = l_Lean_throwAlreadyImported___rarg___closed__1; +x_15 = lean_string_append(x_14, x_13); +lean_dec(x_13); +x_16 = l_Lean_throwAlreadyImported___rarg___closed__2; +x_17 = lean_string_append(x_15, x_16); +x_18 = lean_string_append(x_17, x_11); +lean_dec(x_11); +x_19 = l_Lean_throwAlreadyImported___rarg___closed__3; +x_20 = lean_string_append(x_18, x_19); +x_21 = !lean_is_exclusive(x_2); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; size_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_22 = lean_ctor_get(x_2, 1); +x_23 = lean_ctor_get(x_2, 0); +lean_dec(x_23); +x_24 = lean_array_get_size(x_22); +x_25 = l_Lean_Name_hash___override(x_4); +x_26 = 32; +x_27 = lean_uint64_shift_right(x_25, x_26); +x_28 = lean_uint64_xor(x_25, x_27); +x_29 = 16; +x_30 = lean_uint64_shift_right(x_28, x_29); +x_31 = lean_uint64_xor(x_28, x_30); +x_32 = lean_uint64_to_usize(x_31); +x_33 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_34 = 1; +x_35 = lean_usize_sub(x_33, x_34); +x_36 = lean_usize_land(x_32, x_35); +x_37 = lean_array_uget(x_22, x_36); +lean_dec(x_22); +x_38 = l_Lean_instInhabitedModuleIdx; +x_39 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(x_38, x_4, x_37); +lean_dec(x_37); +lean_dec(x_4); +x_40 = lean_nat_dec_lt(x_39, x_7); +lean_dec(x_7); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_39); +x_41 = l_Lean_instInhabitedName; +x_42 = l_outOfBounds___rarg(x_41); +x_43 = l_Lean_Name_toString(x_42, x_9, x_10); +x_44 = lean_string_append(x_20, x_43); +lean_dec(x_43); +x_45 = l_Lean_instToStringImport___closed__2; +x_46 = lean_string_append(x_44, x_45); +x_47 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 1, x_5); +lean_ctor_set(x_2, 0, x_47); +return x_2; } -if (lean_is_scalar(x_217)) { - x_218 = lean_alloc_ctor(1, 2, 0); -} else { - x_218 = x_217; +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_48 = lean_array_fget(x_6, x_39); +lean_dec(x_39); +x_49 = l_Lean_Name_toString(x_48, x_9, x_10); +x_50 = lean_string_append(x_20, x_49); +lean_dec(x_49); +x_51 = l_Lean_instToStringImport___closed__2; +x_52 = lean_string_append(x_50, x_51); +x_53 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 1, x_5); +lean_ctor_set(x_2, 0, x_53); +return x_2; } -lean_ctor_set(x_218, 0, x_215); -lean_ctor_set(x_218, 1, x_216); -return x_218; } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; size_t x_224; size_t x_225; -x_219 = lean_box(0); -lean_inc(x_2); -x_220 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_201, x_189, x_203, x_219, x_9); -x_221 = lean_ctor_get(x_220, 0); -lean_inc(x_221); -x_222 = lean_ctor_get(x_220, 1); -lean_inc(x_222); -lean_dec(x_220); -x_223 = lean_ctor_get(x_221, 0); -lean_inc(x_223); -lean_dec(x_221); -x_224 = 1; -x_225 = lean_usize_add(x_7, x_224); -x_7 = x_225; -x_8 = x_223; -x_9 = x_222; -goto _start; +lean_object* x_54; lean_object* x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; size_t x_63; size_t x_64; size_t x_65; size_t x_66; size_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_54 = lean_ctor_get(x_2, 1); +lean_inc(x_54); +lean_dec(x_2); +x_55 = lean_array_get_size(x_54); +x_56 = l_Lean_Name_hash___override(x_4); +x_57 = 32; +x_58 = lean_uint64_shift_right(x_56, x_57); +x_59 = lean_uint64_xor(x_56, x_58); +x_60 = 16; +x_61 = lean_uint64_shift_right(x_59, x_60); +x_62 = lean_uint64_xor(x_59, x_61); +x_63 = lean_uint64_to_usize(x_62); +x_64 = lean_usize_of_nat(x_55); +lean_dec(x_55); +x_65 = 1; +x_66 = lean_usize_sub(x_64, x_65); +x_67 = lean_usize_land(x_63, x_66); +x_68 = lean_array_uget(x_54, x_67); +lean_dec(x_54); +x_69 = l_Lean_instInhabitedModuleIdx; +x_70 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(x_69, x_4, x_68); +lean_dec(x_68); +lean_dec(x_4); +x_71 = lean_nat_dec_lt(x_70, x_7); +lean_dec(x_7); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_70); +x_72 = l_Lean_instInhabitedName; +x_73 = l_outOfBounds___rarg(x_72); +x_74 = l_Lean_Name_toString(x_73, x_9, x_10); +x_75 = lean_string_append(x_20, x_74); +lean_dec(x_74); +x_76 = l_Lean_instToStringImport___closed__2; +x_77 = lean_string_append(x_75, x_76); +x_78 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_78, 0, x_77); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_5); +return x_79; } +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; +x_80 = lean_array_fget(x_6, x_70); +lean_dec(x_70); +x_81 = l_Lean_Name_toString(x_80, x_9, x_10); +x_82 = lean_string_append(x_20, x_81); +lean_dec(x_81); +x_83 = l_Lean_instToStringImport___closed__2; +x_84 = lean_string_append(x_82, x_83); +x_85 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_5); +return x_86; } } } } } -else +LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported(lean_object* x_1) { +_start: { -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; uint8_t x_271; -x_263 = lean_ctor_get(x_8, 1); -x_264 = lean_ctor_get(x_8, 0); -lean_inc(x_263); -lean_inc(x_264); -lean_dec(x_8); -x_265 = lean_ctor_get(x_263, 0); -lean_inc(x_265); -x_266 = lean_ctor_get(x_263, 1); -lean_inc(x_266); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_267 = x_263; -} else { - lean_dec_ref(x_263); - x_267 = lean_box(0); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_throwAlreadyImported___rarg___boxed), 5, 0); +return x_2; } -x_268 = lean_ctor_get(x_264, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_264, 1); -lean_inc(x_269); -x_270 = lean_ctor_get(x_264, 2); -lean_inc(x_270); -x_271 = lean_nat_dec_lt(x_269, x_270); -if (x_271 == 0) +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_272; lean_object* x_273; lean_object* x_274; -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_12); +lean_object* x_4; +x_4 = l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_throwAlreadyImported___spec__1(x_1, x_2, x_3); +lean_dec(x_3); lean_dec(x_2); -if (lean_is_scalar(x_267)) { - x_272 = lean_alloc_ctor(0, 2, 0); -} else { - x_272 = x_267; +return x_4; } -lean_ctor_set(x_272, 0, x_265); -lean_ctor_set(x_272, 1, x_266); -x_273 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_273, 0, x_264); -lean_ctor_set(x_273, 1, x_272); -x_274 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_274, 0, x_273); -lean_ctor_set(x_274, 1, x_9); -return x_274; +} +LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_throwAlreadyImported___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_ImportStateM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_st_mk_ref(x_2, x_3); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_8 = lean_apply_2(x_1, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_st_ref_get(x_6, x_10); +lean_dec(x_6); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_11, 0); +lean_ctor_set(x_4, 1, x_13); +lean_ctor_set(x_4, 0, x_9); +lean_ctor_set(x_11, 0, x_4); +return x_11; } else { -lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; uint64_t x_310; uint64_t x_311; uint64_t x_312; uint64_t x_313; uint64_t x_314; uint64_t x_315; uint64_t x_316; size_t x_317; size_t x_318; size_t x_319; size_t x_320; size_t x_321; lean_object* x_322; lean_object* x_323; -lean_dec(x_267); -if (lean_is_exclusive(x_264)) { - lean_ctor_release(x_264, 0); - lean_ctor_release(x_264, 1); - lean_ctor_release(x_264, 2); - x_275 = x_264; -} else { - lean_dec_ref(x_264); - x_275 = lean_box(0); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_11); +lean_ctor_set(x_4, 1, x_14); +lean_ctor_set(x_4, 0, x_9); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_4); +lean_ctor_set(x_16, 1, x_15); +return x_16; } -x_276 = lean_array_fget(x_268, x_269); -x_277 = lean_unsigned_to_nat(1u); -x_278 = lean_nat_add(x_269, x_277); -lean_dec(x_269); -if (lean_is_scalar(x_275)) { - x_279 = lean_alloc_ctor(0, 3, 0); -} else { - x_279 = x_275; } -lean_ctor_set(x_279, 0, x_268); -lean_ctor_set(x_279, 1, x_278); -lean_ctor_set(x_279, 2, x_270); -x_306 = lean_ctor_get(x_266, 0); -lean_inc(x_306); -x_307 = lean_ctor_get(x_266, 1); -lean_inc(x_307); -if (lean_is_exclusive(x_266)) { - lean_ctor_release(x_266, 0); - lean_ctor_release(x_266, 1); - x_308 = x_266; -} else { - lean_dec_ref(x_266); - x_308 = lean_box(0); +else +{ +uint8_t x_17; +lean_free_object(x_4); +lean_dec(x_6); +x_17 = !lean_is_exclusive(x_8); +if (x_17 == 0) +{ +return x_8; } -x_309 = lean_array_get_size(x_307); -x_310 = l_Lean_Name_hash___override(x_12); -x_311 = 32; -x_312 = lean_uint64_shift_right(x_310, x_311); -x_313 = lean_uint64_xor(x_310, x_312); -x_314 = 16; -x_315 = lean_uint64_shift_right(x_313, x_314); -x_316 = lean_uint64_xor(x_313, x_315); -x_317 = lean_uint64_to_usize(x_316); -x_318 = lean_usize_of_nat(x_309); -lean_dec(x_309); -x_319 = 1; -x_320 = lean_usize_sub(x_318, x_319); -x_321 = lean_usize_land(x_317, x_320); -x_322 = lean_array_uget(x_307, x_321); -x_323 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_12, x_322); -if (lean_obj_tag(x_323) == 0) +else { -lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; uint8_t x_333; -x_324 = lean_nat_add(x_306, x_277); -lean_dec(x_306); -lean_inc(x_276); -lean_inc(x_12); -x_325 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_325, 0, x_12); -lean_ctor_set(x_325, 1, x_276); -lean_ctor_set(x_325, 2, x_322); -x_326 = lean_array_uset(x_307, x_321, x_325); -x_327 = lean_box(0); -x_328 = lean_unsigned_to_nat(4u); -x_329 = lean_nat_mul(x_324, x_328); -x_330 = lean_unsigned_to_nat(3u); -x_331 = lean_nat_div(x_329, x_330); -lean_dec(x_329); -x_332 = lean_array_get_size(x_326); -x_333 = lean_nat_dec_le(x_331, x_332); -lean_dec(x_332); -lean_dec(x_331); -if (x_333 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_8); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +else { -lean_object* x_334; lean_object* x_335; -x_334 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_326); -if (lean_is_scalar(x_308)) { - x_335 = lean_alloc_ctor(0, 2, 0); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_4, 0); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_4); +lean_inc(x_21); +x_23 = lean_apply_2(x_1, x_21, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_st_ref_get(x_21, x_25); +lean_dec(x_21); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_29 = x_26; } else { - x_335 = x_308; + lean_dec_ref(x_26); + x_29 = lean_box(0); } -lean_ctor_set(x_335, 0, x_324); -lean_ctor_set(x_335, 1, x_334); -x_280 = x_327; -x_281 = x_335; -goto block_305; +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_24); +lean_ctor_set(x_30, 1, x_27); +if (lean_is_scalar(x_29)) { + x_31 = lean_alloc_ctor(0, 2, 0); +} else { + x_31 = x_29; +} +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_21); +x_32 = lean_ctor_get(x_23, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_23, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + lean_ctor_release(x_23, 1); + x_34 = x_23; +} else { + lean_dec_ref(x_23); + x_34 = lean_box(0); +} +if (lean_is_scalar(x_34)) { + x_35 = lean_alloc_ctor(1, 2, 0); +} else { + x_35 = x_34; +} +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_ImportStateM_run(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_ImportStateM_run___rarg), 3, 0); +return x_2; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_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_6; +x_6 = lean_read_module_data(x_1, x_5); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +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); +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 = 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); +lean_dec(x_14); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; 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_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +x_27 = l_Array_forIn_x27Unsafe_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_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_forIn_x27Unsafe_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_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_39); + x_41 = lean_box(0); +} +x_42 = l_Array_forIn_x27Unsafe_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_43 = x_41; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; +} +} +else +{ +uint8_t x_44; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +x_44 = !lean_is_exclusive(x_12); +if (x_44 == 0) +{ +return x_12; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_12, 0); +x_46 = lean_ctor_get(x_12, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_12); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +uint8_t x_48; +lean_dec(x_2); +x_48 = !lean_is_exclusive(x_6); +if (x_48 == 0) +{ +return x_6; +} +else +{ +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; +} +} +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("object file '", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("' of module ", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_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; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_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_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_10); +x_11 = l_Lean_NameHashSet_insert(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); +x_14 = l_Lean_findOLean(x_10, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_ctor_get(x_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) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = l_Array_forIn_x27Unsafe_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_forIn_x27Unsafe_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_instToStringImport___closed__1; +x_28 = l_Lean_Name_toString(x_10, x_26, x_27); +x_29 = lean_string_append(x_25, x_28); +lean_dec(x_28); +x_30 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___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_17, 1); +lean_ctor_set(x_17, 0, x_32); +return x_17; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_33 = lean_ctor_get(x_17, 1); +lean_inc(x_33); +lean_dec(x_17); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; +x_35 = lean_string_append(x_34, x_15); +lean_dec(x_15); +x_36 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; +x_37 = lean_string_append(x_35, x_36); +x_38 = 1; +x_39 = l_Lean_instToStringImport___closed__1; +x_40 = l_Lean_Name_toString(x_10, x_38, x_39); +x_41 = lean_string_append(x_37, x_40); +lean_dec(x_40); +x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; +x_43 = lean_string_append(x_41, x_42); +x_44 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_33); +return x_45; +} +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_17, 1); +lean_inc(x_46); +lean_dec(x_17); +x_47 = lean_box(0); +x_48 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_15, x_10, x_47, x_3, x_46); +lean_dec(x_15); +return x_48; +} +} +else +{ +uint8_t x_49; +lean_dec(x_10); +x_49 = !lean_is_exclusive(x_14); +if (x_49 == 0) +{ +return x_14; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_14, 0); +x_51 = lean_ctor_get(x_14, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_14); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_53 = lean_ctor_get(x_6, 0); +x_54 = lean_ctor_get(x_6, 1); +x_55 = lean_ctor_get(x_6, 2); +x_56 = lean_ctor_get(x_6, 3); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_6); +x_57 = lean_ctor_get(x_1, 0); +lean_inc(x_57); +lean_dec(x_1); +lean_inc(x_57); +x_58 = l_Lean_NameHashSet_insert(x_53, x_57); +x_59 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_54); +lean_ctor_set(x_59, 2, x_55); +lean_ctor_set(x_59, 3, x_56); +x_60 = lean_st_ref_set(x_3, x_59, x_7); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = l_Lean_findOLean(x_57, x_61); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = l_System_FilePath_pathExists(x_63, x_64); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_unbox(x_66); +lean_dec(x_66); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_69 = x_65; +} else { + lean_dec_ref(x_65); + x_69 = lean_box(0); +} +x_70 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__1; +x_71 = lean_string_append(x_70, x_63); +lean_dec(x_63); +x_72 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__2; +x_73 = lean_string_append(x_71, x_72); +x_74 = 1; +x_75 = l_Lean_instToStringImport___closed__1; +x_76 = l_Lean_Name_toString(x_57, x_74, x_75); +x_77 = lean_string_append(x_73, x_76); +lean_dec(x_76); +x_78 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; +x_79 = lean_string_append(x_77, x_78); +x_80 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_80, 0, x_79); +if (lean_is_scalar(x_69)) { + x_81 = lean_alloc_ctor(1, 2, 0); +} else { + x_81 = x_69; + lean_ctor_set_tag(x_81, 1); +} +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_68); +return x_81; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_65, 1); +lean_inc(x_82); +lean_dec(x_65); +x_83 = lean_box(0); +x_84 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__1(x_63, x_57, x_83, x_3, x_82); +lean_dec(x_63); +return x_84; +} +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_57); +x_85 = lean_ctor_get(x_62, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_62, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_87 = x_62; +} else { + lean_dec_ref(x_62); + 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; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_5, x_4); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +lean_dec(x_6); +x_11 = lean_array_uget(x_3, x_5); +x_12 = lean_st_ref_get(x_7, x_8); +x_13 = lean_ctor_get_uint8(x_11, sizeof(void*)*1); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_ctor_get(x_11, 0); +lean_inc(x_17); +x_18 = l_Lean_NameHashSet_contains(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +x_20 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2(x_11, x_19, x_7, x_15); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +if (lean_obj_tag(x_21) == 0) +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 0); +lean_dec(x_23); +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec(x_21); +lean_ctor_set(x_20, 0, x_24); +return x_20; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_20, 1); +lean_inc(x_25); +lean_dec(x_20); +x_26 = lean_ctor_get(x_21, 0); +lean_inc(x_26); +lean_dec(x_21); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; +x_28 = lean_ctor_get(x_20, 1); +lean_inc(x_28); +lean_dec(x_20); +x_29 = lean_ctor_get(x_21, 0); +lean_inc(x_29); +lean_dec(x_21); +x_30 = 1; +x_31 = lean_usize_add(x_5, x_30); +x_5 = x_31; +x_6 = x_29; +x_8 = x_28; +goto _start; +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_20); +if (x_33 == 0) +{ +return x_20; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_20, 0); +x_35 = lean_ctor_get(x_20, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_20); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +else +{ +size_t x_37; size_t x_38; lean_object* x_39; +lean_dec(x_11); +x_37 = 1; +x_38 = lean_usize_add(x_5, x_37); +x_39 = lean_box(0); +x_5 = x_38; +x_6 = x_39; +x_8 = x_15; +goto _start; +} +} +else +{ +lean_object* x_41; size_t x_42; size_t x_43; lean_object* x_44; +lean_dec(x_11); +x_41 = lean_ctor_get(x_12, 1); +lean_inc(x_41); +lean_dec(x_12); +x_42 = 1; +x_43 = lean_usize_add(x_5, x_42); +x_44 = lean_box(0); +x_5 = x_43; +x_6 = x_44; +x_8 = x_41; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_importModulesCore(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; size_t x_5; size_t x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_box(0); +x_5 = lean_array_size(x_1); +x_6 = 0; +x_7 = lean_box(0); +x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1(x_1, x_4, x_1, x_5, x_6, x_7, x_2, x_3); +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; +x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_10); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +else +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_7); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_8); +if (x_13 == 0) +{ +return x_8; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_8, 0); +x_15 = lean_ctor_get(x_8, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_8); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_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_forIn_x27Unsafe_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_forIn_x27Unsafe_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_forIn_x27Unsafe_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; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___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_4); +lean_dec(x_4); +x_10 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_11 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModulesCore___spec__1(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_importModulesCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_importModulesCore(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Environment_0__Lean_equivInfo(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 2) +{ +if (lean_obj_tag(x_2) == 2) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_5, 0); +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_7, 0); +x_9 = lean_name_eq(x_6, x_8); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_5, 2); +x_12 = lean_ctor_get(x_7, 2); +x_13 = lean_expr_eqv(x_11, x_12); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 0; +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_ctor_get(x_7, 1); +x_17 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_431____spec__1(x_15, x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_3, 2); +x_20 = lean_ctor_get(x_4, 2); +x_21 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_431____spec__1(x_19, x_20); +return x_21; +} +} +} +} +else +{ +uint8_t x_22; +x_22 = 0; +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = 0; +return x_23; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_equivInfo___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Environment_0__Lean_equivInfo(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_finalizeImport_unsafe__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_runtime_mark_persistent(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_finalizeImport_unsafe__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_runtime_mark_persistent(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_name_eq(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_finalizeImport___spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = l_Lean_Name_hash___override(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = l_Lean_Name_hash___override(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_finalizeImport___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_finalizeImport___spec__4(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_finalizeImport___spec__3(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___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) { +_start: +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; size_t x_19; size_t x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; uint8_t x_25; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_array_get_size(x_10); +x_12 = l_Lean_Name_hash___override(x_1); +x_13 = 32; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = 16; +x_17 = lean_uint64_shift_right(x_15, x_16); +x_18 = lean_uint64_xor(x_15, x_17); +x_19 = lean_uint64_to_usize(x_18); +x_20 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_21 = 1; +x_22 = lean_usize_sub(x_20, x_21); +x_23 = lean_usize_land(x_19, x_22); +x_24 = lean_array_uget(x_10, x_23); +x_25 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_1, x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_9, x_26); +lean_dec(x_9); +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_1); +lean_ctor_set(x_28, 1, x_2); +lean_ctor_set(x_28, 2, x_24); +x_29 = lean_array_uset(x_10, x_23, x_28); +x_30 = lean_unsigned_to_nat(4u); +x_31 = lean_nat_mul(x_27, x_30); +x_32 = lean_unsigned_to_nat(3u); +x_33 = lean_nat_div(x_31, x_32); +lean_dec(x_31); +x_34 = lean_array_get_size(x_29); +x_35 = lean_nat_dec_le(x_33, x_34); +lean_dec(x_34); +lean_dec(x_33); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_36 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(x_29); +lean_ctor_set(x_4, 1, x_36); +lean_ctor_set(x_4, 0, x_27); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_4); +lean_ctor_set(x_37, 1, x_5); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_3); +lean_ctor_set(x_38, 1, x_37); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_7); +return x_40; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_ctor_set(x_4, 1, x_29); +lean_ctor_set(x_4, 0, x_27); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_4); +lean_ctor_set(x_41, 1, x_5); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_3); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_7); +return x_44; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_24); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_4); +lean_ctor_set(x_45, 1, x_5); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_3); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_47, 0, x_46); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_7); +return x_48; +} +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; size_t x_59; size_t x_60; size_t x_61; size_t x_62; size_t x_63; lean_object* x_64; uint8_t x_65; +x_49 = lean_ctor_get(x_4, 0); +x_50 = lean_ctor_get(x_4, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_4); +x_51 = lean_array_get_size(x_50); +x_52 = l_Lean_Name_hash___override(x_1); +x_53 = 32; +x_54 = lean_uint64_shift_right(x_52, x_53); +x_55 = lean_uint64_xor(x_52, x_54); +x_56 = 16; +x_57 = lean_uint64_shift_right(x_55, x_56); +x_58 = lean_uint64_xor(x_55, x_57); +x_59 = lean_uint64_to_usize(x_58); +x_60 = lean_usize_of_nat(x_51); +lean_dec(x_51); +x_61 = 1; +x_62 = lean_usize_sub(x_60, x_61); +x_63 = lean_usize_land(x_59, x_62); +x_64 = lean_array_uget(x_50, x_63); +x_65 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_1, 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; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_66 = lean_unsigned_to_nat(1u); +x_67 = lean_nat_add(x_49, x_66); +lean_dec(x_49); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_1); +lean_ctor_set(x_68, 1, x_2); +lean_ctor_set(x_68, 2, x_64); +x_69 = lean_array_uset(x_50, x_63, x_68); +x_70 = lean_unsigned_to_nat(4u); +x_71 = lean_nat_mul(x_67, x_70); +x_72 = lean_unsigned_to_nat(3u); +x_73 = lean_nat_div(x_71, x_72); +lean_dec(x_71); +x_74 = lean_array_get_size(x_69); +x_75 = lean_nat_dec_le(x_73, x_74); +lean_dec(x_74); +lean_dec(x_73); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_76 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(x_69); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_67); +lean_ctor_set(x_77, 1, x_76); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_5); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_3); +lean_ctor_set(x_79, 1, x_78); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_7); +return x_81; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_67); +lean_ctor_set(x_82, 1, x_69); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_5); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_3); +lean_ctor_set(x_84, 1, x_83); +x_85 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_85, 0, 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_7); +return x_86; +} +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_dec(x_64); +lean_dec(x_2); +lean_dec(x_1); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_49); +lean_ctor_set(x_87, 1, x_50); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_5); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_3); +lean_ctor_set(x_89, 1, x_88); +x_90 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_90, 0, x_89); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_7); +return x_91; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_7, x_6); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_2); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +else +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_array_uget(x_5, x_7); +x_13 = !lean_is_exclusive(x_8); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_8, 1); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_14, 0); +x_18 = lean_ctor_get(x_14, 1); +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +x_21 = lean_ctor_get(x_16, 2); +lean_inc(x_21); +x_22 = lean_nat_dec_lt(x_20, x_21); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_12); +lean_dec(x_2); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_8); +lean_ctor_set(x_23, 1, x_9); +return x_23; +} +else +{ +uint8_t x_24; +lean_free_object(x_14); +lean_free_object(x_8); +x_24 = !lean_is_exclusive(x_16); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_57; +x_25 = lean_ctor_get(x_16, 2); +lean_dec(x_25); +x_26 = lean_ctor_get(x_16, 1); +lean_dec(x_26); +x_27 = lean_ctor_get(x_16, 0); +lean_dec(x_27); +x_28 = lean_array_fget(x_19, x_20); +x_29 = lean_unsigned_to_nat(1u); +x_30 = lean_nat_add(x_20, x_29); +lean_dec(x_20); +lean_ctor_set(x_16, 1, x_30); +x_57 = !lean_is_exclusive(x_18); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint64_t x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; size_t x_68; size_t x_69; size_t x_70; size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; +x_58 = lean_ctor_get(x_18, 0); +x_59 = lean_ctor_get(x_18, 1); +x_60 = lean_array_get_size(x_59); +x_61 = l_Lean_Name_hash___override(x_12); +x_62 = 32; +x_63 = lean_uint64_shift_right(x_61, x_62); +x_64 = lean_uint64_xor(x_61, x_63); +x_65 = 16; +x_66 = lean_uint64_shift_right(x_64, x_65); +x_67 = lean_uint64_xor(x_64, x_66); +x_68 = lean_uint64_to_usize(x_67); +x_69 = lean_usize_of_nat(x_60); +lean_dec(x_60); +x_70 = 1; +x_71 = lean_usize_sub(x_69, x_70); +x_72 = lean_usize_land(x_68, x_71); +x_73 = lean_array_uget(x_59, x_72); +x_74 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_12, x_73); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_75 = lean_nat_add(x_58, x_29); +lean_dec(x_58); +lean_inc(x_28); +lean_inc(x_12); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_12); +lean_ctor_set(x_76, 1, x_28); +lean_ctor_set(x_76, 2, x_73); +x_77 = lean_array_uset(x_59, x_72, x_76); +x_78 = lean_box(0); +x_79 = lean_unsigned_to_nat(4u); +x_80 = lean_nat_mul(x_75, x_79); +x_81 = lean_unsigned_to_nat(3u); +x_82 = lean_nat_div(x_80, x_81); +lean_dec(x_80); +x_83 = lean_array_get_size(x_77); +x_84 = lean_nat_dec_le(x_82, x_83); +lean_dec(x_83); +lean_dec(x_82); +if (x_84 == 0) +{ +lean_object* x_85; +x_85 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_77); +lean_ctor_set(x_18, 1, x_85); +lean_ctor_set(x_18, 0, x_75); +x_31 = x_78; +x_32 = x_18; +goto block_56; } else { -lean_object* x_336; -if (lean_is_scalar(x_308)) { - x_336 = lean_alloc_ctor(0, 2, 0); -} else { - x_336 = x_308; -} -lean_ctor_set(x_336, 0, x_324); -lean_ctor_set(x_336, 1, x_326); -x_280 = x_327; -x_281 = x_336; -goto block_305; +lean_ctor_set(x_18, 1, x_77); +lean_ctor_set(x_18, 0, x_75); +x_31 = x_78; +x_32 = x_18; +goto block_56; } } else { -lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; -lean_dec(x_322); -x_337 = lean_ctor_get(x_323, 0); -lean_inc(x_337); -if (lean_is_exclusive(x_323)) { - lean_ctor_release(x_323, 0); - x_338 = x_323; -} else { - lean_dec_ref(x_323); - x_338 = lean_box(0); +uint8_t x_86; +lean_dec(x_73); +x_86 = !lean_is_exclusive(x_74); +if (x_86 == 0) +{ +x_31 = x_74; +x_32 = x_18; +goto block_56; } -if (lean_is_scalar(x_338)) { - x_339 = lean_alloc_ctor(1, 1, 0); -} else { - x_339 = x_338; +else +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_74, 0); +lean_inc(x_87); +lean_dec(x_74); +x_88 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_88, 0, x_87); +x_31 = x_88; +x_32 = x_18; +goto block_56; } -lean_ctor_set(x_339, 0, x_337); -if (lean_is_scalar(x_308)) { - x_340 = lean_alloc_ctor(0, 2, 0); -} else { - x_340 = x_308; } -lean_ctor_set(x_340, 0, x_306); -lean_ctor_set(x_340, 1, x_307); -x_280 = x_339; -x_281 = x_340; -goto block_305; } -block_305: +else { -if (lean_obj_tag(x_280) == 0) +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint64_t x_92; uint64_t x_93; uint64_t x_94; uint64_t x_95; uint64_t x_96; uint64_t x_97; uint64_t x_98; size_t x_99; size_t x_100; size_t x_101; size_t x_102; size_t x_103; lean_object* x_104; lean_object* x_105; +x_89 = lean_ctor_get(x_18, 0); +x_90 = lean_ctor_get(x_18, 1); +lean_inc(x_90); +lean_inc(x_89); +lean_dec(x_18); +x_91 = lean_array_get_size(x_90); +x_92 = l_Lean_Name_hash___override(x_12); +x_93 = 32; +x_94 = lean_uint64_shift_right(x_92, x_93); +x_95 = lean_uint64_xor(x_92, x_94); +x_96 = 16; +x_97 = lean_uint64_shift_right(x_95, x_96); +x_98 = lean_uint64_xor(x_95, x_97); +x_99 = lean_uint64_to_usize(x_98); +x_100 = lean_usize_of_nat(x_91); +lean_dec(x_91); +x_101 = 1; +x_102 = lean_usize_sub(x_100, x_101); +x_103 = lean_usize_land(x_99, x_102); +x_104 = lean_array_uget(x_90, x_103); +x_105 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_12, x_104); +if (lean_obj_tag(x_105) == 0) { -lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; size_t x_287; size_t x_288; -lean_dec(x_276); -x_282 = lean_box(0); -lean_inc(x_2); -x_283 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_279, x_265, x_281, x_282, x_9); -x_284 = lean_ctor_get(x_283, 0); -lean_inc(x_284); -x_285 = lean_ctor_get(x_283, 1); -lean_inc(x_285); -lean_dec(x_283); -x_286 = lean_ctor_get(x_284, 0); -lean_inc(x_286); -lean_dec(x_284); -x_287 = 1; -x_288 = lean_usize_add(x_7, x_287); -x_7 = x_288; -x_8 = x_286; -x_9 = x_285; -goto _start; +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_106 = lean_nat_add(x_89, x_29); +lean_dec(x_89); +lean_inc(x_28); +lean_inc(x_12); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_12); +lean_ctor_set(x_107, 1, x_28); +lean_ctor_set(x_107, 2, x_104); +x_108 = lean_array_uset(x_90, x_103, x_107); +x_109 = lean_box(0); +x_110 = lean_unsigned_to_nat(4u); +x_111 = lean_nat_mul(x_106, x_110); +x_112 = lean_unsigned_to_nat(3u); +x_113 = lean_nat_div(x_111, x_112); +lean_dec(x_111); +x_114 = lean_array_get_size(x_108); +x_115 = lean_nat_dec_le(x_113, x_114); +lean_dec(x_114); +lean_dec(x_113); +if (x_115 == 0) +{ +lean_object* x_116; lean_object* x_117; +x_116 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_108); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_106); +lean_ctor_set(x_117, 1, x_116); +x_31 = x_109; +x_32 = x_117; +goto block_56; } else { -lean_object* x_290; uint8_t x_291; -x_290 = lean_ctor_get(x_280, 0); -lean_inc(x_290); -lean_dec(x_280); -x_291 = l___private_Lean_Environment_0__Lean_equivInfo(x_290, x_276); -lean_dec(x_276); -lean_dec(x_290); -if (x_291 == 0) +lean_object* x_118; +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_106); +lean_ctor_set(x_118, 1, x_108); +x_31 = x_109; +x_32 = x_118; +goto block_56; +} +} +else { -lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; -lean_dec(x_281); -lean_dec(x_279); -x_292 = l_Lean_throwAlreadyImported___rarg(x_1, x_265, x_2, x_12, x_9); -lean_dec(x_2); -x_293 = lean_ctor_get(x_292, 0); -lean_inc(x_293); -x_294 = lean_ctor_get(x_292, 1); -lean_inc(x_294); -if (lean_is_exclusive(x_292)) { - lean_ctor_release(x_292, 0); - lean_ctor_release(x_292, 1); - x_295 = x_292; +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec(x_104); +x_119 = lean_ctor_get(x_105, 0); +lean_inc(x_119); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + x_120 = x_105; } else { - lean_dec_ref(x_292); - x_295 = lean_box(0); + lean_dec_ref(x_105); + x_120 = lean_box(0); } -if (lean_is_scalar(x_295)) { - x_296 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 1, 0); } else { - x_296 = x_295; + x_121 = x_120; } -lean_ctor_set(x_296, 0, x_293); -lean_ctor_set(x_296, 1, x_294); -return x_296; +lean_ctor_set(x_121, 0, x_119); +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_89); +lean_ctor_set(x_122, 1, x_90); +x_31 = x_121; +x_32 = x_122; +goto block_56; } -else +} +block_56: { -lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; size_t x_302; size_t x_303; -x_297 = lean_box(0); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; size_t x_39; +lean_dec(x_28); +x_33 = lean_box(0); lean_inc(x_2); -x_298 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_279, x_265, x_281, x_297, x_9); -x_299 = lean_ctor_get(x_298, 0); -lean_inc(x_299); -x_300 = lean_ctor_get(x_298, 1); -lean_inc(x_300); -lean_dec(x_298); -x_301 = lean_ctor_get(x_299, 0); -lean_inc(x_301); -lean_dec(x_299); -x_302 = 1; -x_303 = lean_usize_add(x_7, x_302); -x_7 = x_303; -x_8 = x_301; -x_9 = x_300; +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_16, x_17, x_32, x_33, x_9); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); +x_38 = 1; +x_39 = lean_usize_add(x_7, x_38); +x_7 = x_39; +x_8 = x_37; +x_9 = x_36; goto _start; } +else +{ +lean_object* x_41; uint8_t x_42; +x_41 = lean_ctor_get(x_31, 0); +lean_inc(x_41); +lean_dec(x_31); +x_42 = l___private_Lean_Environment_0__Lean_equivInfo(x_41, x_28); +lean_dec(x_28); +lean_dec(x_41); +if (x_42 == 0) +{ +lean_object* x_43; uint8_t x_44; +lean_dec(x_32); +lean_dec(x_16); +x_43 = l_Lean_throwAlreadyImported___rarg(x_1, x_17, x_2, x_12, x_9); +lean_dec(x_2); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +return x_43; } +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_43, 0); +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_43); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; +x_48 = lean_box(0); +lean_inc(x_2); +x_49 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_16, x_17, x_32, x_48, x_9); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_ctor_get(x_50, 0); +lean_inc(x_52); +lean_dec(x_50); +x_53 = 1; +x_54 = lean_usize_add(x_7, x_53); +x_7 = x_54; +x_8 = x_52; +x_9 = x_51; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_6, x_5); -if (x_9 == 0) -{ -lean_object* x_10; -lean_dec(x_1); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_7); -lean_ctor_set(x_10, 1, x_8); -return x_10; -} else { -lean_object* x_11; uint8_t x_12; -x_11 = lean_array_uget(x_4, x_6); -x_12 = !lean_is_exclusive(x_7); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; -x_13 = lean_ctor_get(x_7, 0); -x_14 = lean_ctor_get(x_7, 1); -x_15 = lean_array_get_size(x_14); -x_16 = l_Lean_Name_hash___override(x_11); -x_17 = 32; -x_18 = lean_uint64_shift_right(x_16, x_17); -x_19 = lean_uint64_xor(x_16, x_18); -x_20 = 16; -x_21 = lean_uint64_shift_right(x_19, x_20); -x_22 = lean_uint64_xor(x_19, x_21); -x_23 = lean_uint64_to_usize(x_22); -x_24 = lean_usize_of_nat(x_15); -lean_dec(x_15); -x_25 = 1; -x_26 = lean_usize_sub(x_24, x_25); -x_27 = lean_usize_land(x_23, x_26); -x_28 = lean_array_uget(x_14, x_27); -x_29 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_11, x_28); -if (x_29 == 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_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint64_t x_157; uint64_t x_158; uint64_t x_159; uint64_t x_160; uint64_t x_161; uint64_t x_162; uint64_t x_163; size_t x_164; size_t x_165; size_t x_166; size_t x_167; size_t x_168; lean_object* x_169; lean_object* x_170; +lean_dec(x_16); +x_123 = lean_array_fget(x_19, x_20); +x_124 = lean_unsigned_to_nat(1u); +x_125 = lean_nat_add(x_20, x_124); +lean_dec(x_20); +x_126 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_126, 0, x_19); +lean_ctor_set(x_126, 1, x_125); +lean_ctor_set(x_126, 2, x_21); +x_153 = lean_ctor_get(x_18, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_18, 1); +lean_inc(x_154); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + x_155 = x_18; +} else { + lean_dec_ref(x_18); + x_155 = lean_box(0); +} +x_156 = lean_array_get_size(x_154); +x_157 = l_Lean_Name_hash___override(x_12); +x_158 = 32; +x_159 = lean_uint64_shift_right(x_157, x_158); +x_160 = lean_uint64_xor(x_157, x_159); +x_161 = 16; +x_162 = lean_uint64_shift_right(x_160, x_161); +x_163 = lean_uint64_xor(x_160, x_162); +x_164 = lean_uint64_to_usize(x_163); +x_165 = lean_usize_of_nat(x_156); +lean_dec(x_156); +x_166 = 1; +x_167 = lean_usize_sub(x_165, x_166); +x_168 = lean_usize_land(x_164, x_167); +x_169 = lean_array_uget(x_154, x_168); +x_170 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_12, x_169); +if (lean_obj_tag(x_170) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_13, x_30); -lean_dec(x_13); -lean_inc(x_1); -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_11); -lean_ctor_set(x_32, 1, x_1); -lean_ctor_set(x_32, 2, x_28); -x_33 = lean_array_uset(x_14, x_27, x_32); -x_34 = lean_unsigned_to_nat(4u); -x_35 = lean_nat_mul(x_31, x_34); -x_36 = lean_unsigned_to_nat(3u); -x_37 = lean_nat_div(x_35, x_36); -lean_dec(x_35); -x_38 = lean_array_get_size(x_33); -x_39 = lean_nat_dec_le(x_37, x_38); -lean_dec(x_38); -lean_dec(x_37); -if (x_39 == 0) +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; +x_171 = lean_nat_add(x_153, x_124); +lean_dec(x_153); +lean_inc(x_123); +lean_inc(x_12); +x_172 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_172, 0, x_12); +lean_ctor_set(x_172, 1, x_123); +lean_ctor_set(x_172, 2, x_169); +x_173 = lean_array_uset(x_154, x_168, x_172); +x_174 = lean_box(0); +x_175 = lean_unsigned_to_nat(4u); +x_176 = lean_nat_mul(x_171, x_175); +x_177 = lean_unsigned_to_nat(3u); +x_178 = lean_nat_div(x_176, x_177); +lean_dec(x_176); +x_179 = lean_array_get_size(x_173); +x_180 = lean_nat_dec_le(x_178, x_179); +lean_dec(x_179); +lean_dec(x_178); +if (x_180 == 0) { -lean_object* x_40; size_t x_41; -x_40 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(x_33); -lean_ctor_set(x_7, 1, x_40); -lean_ctor_set(x_7, 0, x_31); -x_41 = lean_usize_add(x_6, x_25); -x_6 = x_41; -goto _start; +lean_object* x_181; lean_object* x_182; +x_181 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_173); +if (lean_is_scalar(x_155)) { + x_182 = lean_alloc_ctor(0, 2, 0); +} else { + x_182 = x_155; +} +lean_ctor_set(x_182, 0, x_171); +lean_ctor_set(x_182, 1, x_181); +x_127 = x_174; +x_128 = x_182; +goto block_152; } else { -size_t x_43; -lean_ctor_set(x_7, 1, x_33); -lean_ctor_set(x_7, 0, x_31); -x_43 = lean_usize_add(x_6, x_25); -x_6 = x_43; -goto _start; +lean_object* x_183; +if (lean_is_scalar(x_155)) { + x_183 = lean_alloc_ctor(0, 2, 0); +} else { + x_183 = x_155; +} +lean_ctor_set(x_183, 0, x_171); +lean_ctor_set(x_183, 1, x_173); +x_127 = x_174; +x_128 = x_183; +goto block_152; } } else { -size_t x_45; -lean_dec(x_28); -lean_dec(x_11); -x_45 = lean_usize_add(x_6, x_25); -x_6 = x_45; -goto _start; +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +lean_dec(x_169); +x_184 = lean_ctor_get(x_170, 0); +lean_inc(x_184); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + x_185 = x_170; +} else { + lean_dec_ref(x_170); + x_185 = lean_box(0); } +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 1, 0); +} else { + x_186 = x_185; } -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; uint64_t x_56; size_t x_57; size_t x_58; size_t x_59; size_t x_60; size_t x_61; lean_object* x_62; uint8_t x_63; -x_47 = lean_ctor_get(x_7, 0); -x_48 = lean_ctor_get(x_7, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_7); -x_49 = lean_array_get_size(x_48); -x_50 = l_Lean_Name_hash___override(x_11); -x_51 = 32; -x_52 = lean_uint64_shift_right(x_50, x_51); -x_53 = lean_uint64_xor(x_50, x_52); -x_54 = 16; -x_55 = lean_uint64_shift_right(x_53, x_54); -x_56 = lean_uint64_xor(x_53, x_55); -x_57 = lean_uint64_to_usize(x_56); -x_58 = lean_usize_of_nat(x_49); -lean_dec(x_49); -x_59 = 1; -x_60 = lean_usize_sub(x_58, x_59); -x_61 = lean_usize_land(x_57, x_60); -x_62 = lean_array_uget(x_48, x_61); -x_63 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_11, x_62); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; -x_64 = lean_unsigned_to_nat(1u); -x_65 = lean_nat_add(x_47, x_64); -lean_dec(x_47); -lean_inc(x_1); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_11); -lean_ctor_set(x_66, 1, x_1); -lean_ctor_set(x_66, 2, x_62); -x_67 = lean_array_uset(x_48, x_61, x_66); -x_68 = lean_unsigned_to_nat(4u); -x_69 = lean_nat_mul(x_65, x_68); -x_70 = lean_unsigned_to_nat(3u); -x_71 = lean_nat_div(x_69, x_70); -lean_dec(x_69); -x_72 = lean_array_get_size(x_67); -x_73 = lean_nat_dec_le(x_71, x_72); -lean_dec(x_72); -lean_dec(x_71); -if (x_73 == 0) +lean_ctor_set(x_186, 0, x_184); +if (lean_is_scalar(x_155)) { + x_187 = lean_alloc_ctor(0, 2, 0); +} else { + x_187 = x_155; +} +lean_ctor_set(x_187, 0, x_153); +lean_ctor_set(x_187, 1, x_154); +x_127 = x_186; +x_128 = x_187; +goto block_152; +} +block_152: { -lean_object* x_74; lean_object* x_75; size_t x_76; -x_74 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(x_67); -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_65); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_usize_add(x_6, x_59); -x_6 = x_76; -x_7 = x_75; +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; size_t x_134; size_t x_135; +lean_dec(x_123); +x_129 = lean_box(0); +lean_inc(x_2); +x_130 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_126, x_17, x_128, x_129, x_9); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +lean_dec(x_130); +x_133 = lean_ctor_get(x_131, 0); +lean_inc(x_133); +lean_dec(x_131); +x_134 = 1; +x_135 = lean_usize_add(x_7, x_134); +x_7 = x_135; +x_8 = x_133; +x_9 = x_132; goto _start; } else { -lean_object* x_78; size_t x_79; -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_65); -lean_ctor_set(x_78, 1, x_67); -x_79 = lean_usize_add(x_6, x_59); -x_6 = x_79; -x_7 = x_78; -goto _start; +lean_object* x_137; uint8_t x_138; +x_137 = lean_ctor_get(x_127, 0); +lean_inc(x_137); +lean_dec(x_127); +x_138 = l___private_Lean_Environment_0__Lean_equivInfo(x_137, x_123); +lean_dec(x_123); +lean_dec(x_137); +if (x_138 == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +lean_dec(x_128); +lean_dec(x_126); +x_139 = l_Lean_throwAlreadyImported___rarg(x_1, x_17, x_2, x_12, x_9); +lean_dec(x_2); +x_140 = lean_ctor_get(x_139, 0); +lean_inc(x_140); +x_141 = lean_ctor_get(x_139, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_142 = x_139; +} else { + lean_dec_ref(x_139); + x_142 = lean_box(0); +} +if (lean_is_scalar(x_142)) { + x_143 = lean_alloc_ctor(1, 2, 0); +} else { + x_143 = x_142; } +lean_ctor_set(x_143, 0, x_140); +lean_ctor_set(x_143, 1, x_141); +return x_143; } else { -lean_object* x_81; size_t x_82; -lean_dec(x_62); -lean_dec(x_11); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_47); -lean_ctor_set(x_81, 1, x_48); -x_82 = lean_usize_add(x_6, x_59); -x_6 = x_82; -x_7 = x_81; +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; size_t x_149; size_t x_150; +x_144 = lean_box(0); +lean_inc(x_2); +x_145 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_126, x_17, x_128, x_144, x_9); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +lean_dec(x_145); +x_148 = lean_ctor_get(x_146, 0); +lean_inc(x_148); +lean_dec(x_146); +x_149 = 1; +x_150 = lean_usize_add(x_7, x_149); +x_7 = x_150; +x_8 = x_148; +x_9 = x_147; goto _start; } } } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___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; uint8_t x_11; -x_10 = lean_ctor_get(x_4, 1); -x_11 = lean_nat_dec_lt(x_6, x_10); -if (x_11 == 0) -{ -lean_object* x_12; -lean_dec(x_6); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_9); -return x_12; } else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_5); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; -x_14 = lean_array_fget(x_2, x_6); -x_15 = lean_ctor_get(x_14, 2); -lean_inc(x_15); -x_16 = lean_array_get_size(x_15); -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Array_toSubarray___rarg(x_15, x_17, x_16); -x_19 = lean_ctor_get(x_14, 1); -lean_inc(x_19); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_18); -lean_ctor_set(x_21, 1, x_5); -x_22 = lean_array_size(x_19); -x_23 = 0; -lean_inc(x_6); -x_24 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_6, x_19, x_20, x_19, x_22, x_23, x_21, x_9); -lean_dec(x_19); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = !lean_is_exclusive(x_26); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_29 = lean_ctor_get(x_26, 0); -x_30 = lean_ctor_get(x_14, 3); -lean_inc(x_30); +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; uint8_t x_194; +x_188 = lean_ctor_get(x_8, 0); +x_189 = lean_ctor_get(x_14, 0); +x_190 = lean_ctor_get(x_14, 1); +lean_inc(x_190); +lean_inc(x_189); lean_dec(x_14); -x_31 = lean_array_size(x_30); -lean_inc(x_6); -x_32 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_6, x_20, x_30, x_30, x_31, x_23, x_29, x_27); -lean_dec(x_30); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -lean_ctor_set(x_26, 0, x_33); -x_35 = lean_ctor_get(x_4, 2); -x_36 = lean_nat_add(x_6, x_35); -lean_dec(x_6); -x_5 = x_26; -x_6 = x_36; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_9 = x_34; -goto _start; +x_191 = lean_ctor_get(x_188, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_188, 1); +lean_inc(x_192); +x_193 = lean_ctor_get(x_188, 2); +lean_inc(x_193); +x_194 = lean_nat_dec_lt(x_192, x_193); +if (x_194 == 0) +{ +lean_object* x_195; lean_object* x_196; +lean_dec(x_193); +lean_dec(x_192); +lean_dec(x_191); +lean_dec(x_12); +lean_dec(x_2); +x_195 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_195, 0, x_189); +lean_ctor_set(x_195, 1, x_190); +lean_ctor_set(x_8, 1, x_195); +x_196 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_196, 0, x_8); +lean_ctor_set(x_196, 1, x_9); +return x_196; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_38 = lean_ctor_get(x_26, 0); -x_39 = lean_ctor_get(x_26, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_26); -x_40 = lean_ctor_get(x_14, 3); -lean_inc(x_40); -lean_dec(x_14); -x_41 = lean_array_size(x_40); -lean_inc(x_6); -x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_6, x_20, x_40, x_40, x_41, x_23, x_38, x_27); -lean_dec(x_40); -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_39); -x_46 = lean_ctor_get(x_4, 2); -x_47 = lean_nat_add(x_6, x_46); -lean_dec(x_6); -x_5 = x_45; -x_6 = x_47; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_9 = x_44; -goto _start; +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_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; uint64_t x_232; uint64_t x_233; uint64_t x_234; uint64_t x_235; uint64_t x_236; uint64_t x_237; uint64_t x_238; size_t x_239; size_t x_240; size_t x_241; size_t x_242; size_t x_243; lean_object* x_244; lean_object* x_245; +lean_free_object(x_8); +if (lean_is_exclusive(x_188)) { + lean_ctor_release(x_188, 0); + lean_ctor_release(x_188, 1); + lean_ctor_release(x_188, 2); + x_197 = x_188; +} else { + lean_dec_ref(x_188); + x_197 = lean_box(0); +} +x_198 = lean_array_fget(x_191, x_192); +x_199 = lean_unsigned_to_nat(1u); +x_200 = lean_nat_add(x_192, x_199); +lean_dec(x_192); +if (lean_is_scalar(x_197)) { + x_201 = lean_alloc_ctor(0, 3, 0); +} else { + x_201 = x_197; } +lean_ctor_set(x_201, 0, x_191); +lean_ctor_set(x_201, 1, x_200); +lean_ctor_set(x_201, 2, x_193); +x_228 = lean_ctor_get(x_190, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_190, 1); +lean_inc(x_229); +if (lean_is_exclusive(x_190)) { + lean_ctor_release(x_190, 0); + lean_ctor_release(x_190, 1); + x_230 = x_190; +} else { + lean_dec_ref(x_190); + x_230 = lean_box(0); } -else +x_231 = lean_array_get_size(x_229); +x_232 = l_Lean_Name_hash___override(x_12); +x_233 = 32; +x_234 = lean_uint64_shift_right(x_232, x_233); +x_235 = lean_uint64_xor(x_232, x_234); +x_236 = 16; +x_237 = lean_uint64_shift_right(x_235, x_236); +x_238 = lean_uint64_xor(x_235, x_237); +x_239 = lean_uint64_to_usize(x_238); +x_240 = lean_usize_of_nat(x_231); +lean_dec(x_231); +x_241 = 1; +x_242 = lean_usize_sub(x_240, x_241); +x_243 = lean_usize_land(x_239, x_242); +x_244 = lean_array_uget(x_229, x_243); +x_245 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_12, x_244); +if (lean_obj_tag(x_245) == 0) { -uint8_t x_49; -lean_dec(x_14); -lean_dec(x_6); -x_49 = !lean_is_exclusive(x_24); -if (x_49 == 0) +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; uint8_t x_255; +x_246 = lean_nat_add(x_228, x_199); +lean_dec(x_228); +lean_inc(x_198); +lean_inc(x_12); +x_247 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_247, 0, x_12); +lean_ctor_set(x_247, 1, x_198); +lean_ctor_set(x_247, 2, x_244); +x_248 = lean_array_uset(x_229, x_243, x_247); +x_249 = lean_box(0); +x_250 = lean_unsigned_to_nat(4u); +x_251 = lean_nat_mul(x_246, x_250); +x_252 = lean_unsigned_to_nat(3u); +x_253 = lean_nat_div(x_251, x_252); +lean_dec(x_251); +x_254 = lean_array_get_size(x_248); +x_255 = lean_nat_dec_le(x_253, x_254); +lean_dec(x_254); +lean_dec(x_253); +if (x_255 == 0) { -return x_24; +lean_object* x_256; lean_object* x_257; +x_256 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_248); +if (lean_is_scalar(x_230)) { + x_257 = lean_alloc_ctor(0, 2, 0); +} else { + x_257 = x_230; +} +lean_ctor_set(x_257, 0, x_246); +lean_ctor_set(x_257, 1, x_256); +x_202 = x_249; +x_203 = x_257; +goto block_227; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_24, 0); -x_51 = lean_ctor_get(x_24, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_24); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +lean_object* x_258; +if (lean_is_scalar(x_230)) { + x_258 = lean_alloc_ctor(0, 2, 0); +} else { + x_258 = x_230; } +lean_ctor_set(x_258, 0, x_246); +lean_ctor_set(x_258, 1, x_248); +x_202 = x_249; +x_203 = x_258; +goto block_227; } } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; size_t x_64; size_t x_65; lean_object* x_66; -x_53 = lean_ctor_get(x_5, 0); -x_54 = lean_ctor_get(x_5, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_5); -x_55 = lean_array_fget(x_2, x_6); -x_56 = lean_ctor_get(x_55, 2); -lean_inc(x_56); -x_57 = lean_array_get_size(x_56); -x_58 = lean_unsigned_to_nat(0u); -x_59 = l_Array_toSubarray___rarg(x_56, x_58, x_57); -x_60 = lean_ctor_get(x_55, 1); -lean_inc(x_60); -x_61 = lean_box(0); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_53); -lean_ctor_set(x_62, 1, x_54); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_59); -lean_ctor_set(x_63, 1, x_62); -x_64 = lean_array_size(x_60); -x_65 = 0; -lean_inc(x_6); -x_66 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_6, x_60, x_61, x_60, x_64, x_65, x_63, x_9); -lean_dec(x_60); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -lean_dec(x_67); -x_69 = lean_ctor_get(x_66, 1); -lean_inc(x_69); -lean_dec(x_66); -x_70 = lean_ctor_get(x_68, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_68, 1); -lean_inc(x_71); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_72 = x_68; +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +lean_dec(x_244); +x_259 = lean_ctor_get(x_245, 0); +lean_inc(x_259); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + x_260 = x_245; } else { - lean_dec_ref(x_68); - x_72 = lean_box(0); + lean_dec_ref(x_245); + x_260 = lean_box(0); } -x_73 = lean_ctor_get(x_55, 3); -lean_inc(x_73); -lean_dec(x_55); -x_74 = lean_array_size(x_73); -lean_inc(x_6); -x_75 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_6, x_61, x_73, x_73, x_74, x_65, x_70, x_69); -lean_dec(x_73); -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -lean_dec(x_75); -if (lean_is_scalar(x_72)) { - x_78 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_260)) { + x_261 = lean_alloc_ctor(1, 1, 0); } else { - x_78 = x_72; + x_261 = x_260; } -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_71); -x_79 = lean_ctor_get(x_4, 2); -x_80 = lean_nat_add(x_6, x_79); -lean_dec(x_6); -x_5 = x_78; -x_6 = x_80; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_9 = x_77; +lean_ctor_set(x_261, 0, x_259); +if (lean_is_scalar(x_230)) { + x_262 = lean_alloc_ctor(0, 2, 0); +} else { + x_262 = x_230; +} +lean_ctor_set(x_262, 0, x_228); +lean_ctor_set(x_262, 1, x_229); +x_202 = x_261; +x_203 = x_262; +goto block_227; +} +block_227: +{ +if (lean_obj_tag(x_202) == 0) +{ +lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; size_t x_209; size_t x_210; +lean_dec(x_198); +x_204 = lean_box(0); +lean_inc(x_2); +x_205 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_201, x_189, x_203, x_204, x_9); +x_206 = lean_ctor_get(x_205, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_205, 1); +lean_inc(x_207); +lean_dec(x_205); +x_208 = lean_ctor_get(x_206, 0); +lean_inc(x_208); +lean_dec(x_206); +x_209 = 1; +x_210 = lean_usize_add(x_7, x_209); +x_7 = x_210; +x_8 = x_208; +x_9 = x_207; goto _start; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_55); -lean_dec(x_6); -x_82 = lean_ctor_get(x_66, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_66, 1); -lean_inc(x_83); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_84 = x_66; +lean_object* x_212; uint8_t x_213; +x_212 = lean_ctor_get(x_202, 0); +lean_inc(x_212); +lean_dec(x_202); +x_213 = l___private_Lean_Environment_0__Lean_equivInfo(x_212, x_198); +lean_dec(x_198); +lean_dec(x_212); +if (x_213 == 0) +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +lean_dec(x_203); +lean_dec(x_201); +x_214 = l_Lean_throwAlreadyImported___rarg(x_1, x_189, x_2, x_12, x_9); +lean_dec(x_2); +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_214, 1); +lean_inc(x_216); +if (lean_is_exclusive(x_214)) { + lean_ctor_release(x_214, 0); + lean_ctor_release(x_214, 1); + x_217 = x_214; } else { - lean_dec_ref(x_66); - x_84 = lean_box(0); + lean_dec_ref(x_214); + x_217 = lean_box(0); } -if (lean_is_scalar(x_84)) { - x_85 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_217)) { + x_218 = lean_alloc_ctor(1, 2, 0); } else { - x_85 = x_84; + x_218 = x_217; +} +lean_ctor_set(x_218, 0, x_215); +lean_ctor_set(x_218, 1, x_216); +return x_218; +} +else +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; size_t x_224; size_t x_225; +x_219 = lean_box(0); +lean_inc(x_2); +x_220 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_201, x_189, x_203, x_219, x_9); +x_221 = lean_ctor_get(x_220, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_220, 1); +lean_inc(x_222); +lean_dec(x_220); +x_223 = lean_ctor_get(x_221, 0); +lean_inc(x_223); +lean_dec(x_221); +x_224 = 1; +x_225 = lean_usize_add(x_7, x_224); +x_7 = x_225; +x_8 = x_223; +x_9 = x_222; +goto _start; } -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_83); -return x_85; } } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_2, x_3); -if (x_5 == 0) +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; size_t x_13; size_t x_14; -x_6 = lean_array_uget(x_1, x_2); -x_7 = lean_ctor_get(x_6, 2); -lean_inc(x_7); -x_8 = lean_array_get_size(x_7); -lean_dec(x_7); -x_9 = lean_nat_add(x_4, x_8); +lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; uint8_t x_271; +x_263 = lean_ctor_get(x_8, 1); +x_264 = lean_ctor_get(x_8, 0); +lean_inc(x_263); +lean_inc(x_264); lean_dec(x_8); -lean_dec(x_4); -x_10 = lean_ctor_get(x_6, 3); -lean_inc(x_10); -lean_dec(x_6); -x_11 = lean_array_get_size(x_10); -lean_dec(x_10); -x_12 = lean_nat_add(x_9, x_11); -lean_dec(x_11); -lean_dec(x_9); -x_13 = 1; -x_14 = lean_usize_add(x_2, x_13); -x_2 = x_14; -x_4 = x_12; -goto _start; +x_265 = lean_ctor_get(x_263, 0); +lean_inc(x_265); +x_266 = lean_ctor_get(x_263, 1); +lean_inc(x_266); +if (lean_is_exclusive(x_263)) { + lean_ctor_release(x_263, 0); + lean_ctor_release(x_263, 1); + x_267 = x_263; +} else { + lean_dec_ref(x_263); + x_267 = lean_box(0); } -else +x_268 = lean_ctor_get(x_264, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_264, 1); +lean_inc(x_269); +x_270 = lean_ctor_get(x_264, 2); +lean_inc(x_270); +x_271 = lean_nat_dec_lt(x_269, x_270); +if (x_271 == 0) { -return x_4; -} +lean_object* x_272; lean_object* x_273; lean_object* x_274; +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_12); +lean_dec(x_2); +if (lean_is_scalar(x_267)) { + x_272 = lean_alloc_ctor(0, 2, 0); +} else { + x_272 = x_267; } +lean_ctor_set(x_272, 0, x_265); +lean_ctor_set(x_272, 1, x_266); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_264); +lean_ctor_set(x_273, 1, x_272); +x_274 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_274, 0, x_273); +lean_ctor_set(x_274, 1, x_9); +return x_274; } -LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_3); -return x_4; -} +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; uint64_t x_310; uint64_t x_311; uint64_t x_312; uint64_t x_313; uint64_t x_314; uint64_t x_315; uint64_t x_316; size_t x_317; size_t x_318; size_t x_319; size_t x_320; size_t x_321; lean_object* x_322; lean_object* x_323; +lean_dec(x_267); +if (lean_is_exclusive(x_264)) { + lean_ctor_release(x_264, 0); + lean_ctor_release(x_264, 1); + lean_ctor_release(x_264, 2); + x_275 = x_264; +} else { + lean_dec_ref(x_264); + x_275 = lean_box(0); } -static lean_object* _init_l_Lean_finalizeImport___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_finalizeImport___lambda__1___boxed), 3, 0); -return x_1; +x_276 = lean_array_fget(x_268, x_269); +x_277 = lean_unsigned_to_nat(1u); +x_278 = lean_nat_add(x_269, x_277); +lean_dec(x_269); +if (lean_is_scalar(x_275)) { + x_279 = lean_alloc_ctor(0, 3, 0); +} else { + x_279 = x_275; } +lean_ctor_set(x_279, 0, x_268); +lean_ctor_set(x_279, 1, x_278); +lean_ctor_set(x_279, 2, x_270); +x_306 = lean_ctor_get(x_266, 0); +lean_inc(x_306); +x_307 = lean_ctor_get(x_266, 1); +lean_inc(x_307); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_308 = x_266; +} else { + lean_dec_ref(x_266); + x_308 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(x_1, x_2, x_7, x_4, x_6); -if (lean_obj_tag(x_8) == 0) +x_309 = lean_array_get_size(x_307); +x_310 = l_Lean_Name_hash___override(x_12); +x_311 = 32; +x_312 = lean_uint64_shift_right(x_310, x_311); +x_313 = lean_uint64_xor(x_310, x_312); +x_314 = 16; +x_315 = lean_uint64_shift_right(x_313, x_314); +x_316 = lean_uint64_xor(x_313, x_315); +x_317 = lean_uint64_to_usize(x_316); +x_318 = lean_usize_of_nat(x_309); +lean_dec(x_309); +x_319 = 1; +x_320 = lean_usize_sub(x_318, x_319); +x_321 = lean_usize_land(x_317, x_320); +x_322 = lean_array_uget(x_307, x_321); +x_323 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_12, x_322); +if (lean_obj_tag(x_323) == 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_finalizeImport___lambda__2___closed__1; -if (x_3 == 0) +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; uint8_t x_333; +x_324 = lean_nat_add(x_306, x_277); +lean_dec(x_306); +lean_inc(x_276); +lean_inc(x_12); +x_325 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_325, 0, x_12); +lean_ctor_set(x_325, 1, x_276); +lean_ctor_set(x_325, 2, x_322); +x_326 = lean_array_uset(x_307, x_321, x_325); +x_327 = lean_box(0); +x_328 = lean_unsigned_to_nat(4u); +x_329 = lean_nat_mul(x_324, x_328); +x_330 = lean_unsigned_to_nat(3u); +x_331 = lean_nat_div(x_329, x_330); +lean_dec(x_329); +x_332 = lean_array_get_size(x_326); +x_333 = lean_nat_dec_le(x_331, x_332); +lean_dec(x_332); +lean_dec(x_331); +if (x_333 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_box(0); -x_13 = lean_apply_3(x_11, x_9, x_12, x_10); -return x_13; +lean_object* x_334; lean_object* x_335; +x_334 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_326); +if (lean_is_scalar(x_308)) { + x_335 = lean_alloc_ctor(0, 2, 0); +} else { + x_335 = x_308; +} +lean_ctor_set(x_335, 0, x_324); +lean_ctor_set(x_335, 1, x_334); +x_280 = x_327; +x_281 = x_335; +goto block_305; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_runtime_mark_persistent(x_9, x_10); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_box(0); -x_18 = lean_apply_3(x_11, x_15, x_17, x_16); -return x_18; +lean_object* x_336; +if (lean_is_scalar(x_308)) { + x_336 = lean_alloc_ctor(0, 2, 0); +} else { + x_336 = x_308; } +lean_ctor_set(x_336, 0, x_324); +lean_ctor_set(x_336, 1, x_326); +x_280 = x_327; +x_281 = x_336; +goto block_305; } -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_8); -if (x_19 == 0) -{ -return x_8; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; +lean_dec(x_322); +x_337 = lean_ctor_get(x_323, 0); +lean_inc(x_337); +if (lean_is_exclusive(x_323)) { + lean_ctor_release(x_323, 0); + x_338 = x_323; +} else { + lean_dec_ref(x_323); + x_338 = lean_box(0); } +if (lean_is_scalar(x_338)) { + x_339 = lean_alloc_ctor(1, 1, 0); +} else { + x_339 = x_338; } +lean_ctor_set(x_339, 0, x_337); +if (lean_is_scalar(x_308)) { + x_340 = lean_alloc_ctor(0, 2, 0); +} else { + x_340 = x_308; } +lean_ctor_set(x_340, 0, x_306); +lean_ctor_set(x_340, 1, x_307); +x_280 = x_339; +x_281 = x_340; +goto block_305; } -LEAN_EXPORT lean_object* l_Lean_finalizeImport(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint32_t x_4, uint8_t x_5, lean_object* x_6) { -_start: +block_305: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_7 = lean_ctor_get(x_1, 2); -lean_inc(x_7); -x_8 = lean_array_get_size(x_7); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_nat_dec_lt(x_9, x_8); -x_11 = lean_box(0); -x_12 = lean_unsigned_to_nat(1u); -lean_inc(x_8); -x_13 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_13, 0, x_9); -lean_ctor_set(x_13, 1, x_8); -lean_ctor_set(x_13, 2, x_12); -if (x_10 == 0) +if (lean_obj_tag(x_280) == 0) { -lean_dec(x_8); -x_14 = x_9; -goto block_64; +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; size_t x_287; size_t x_288; +lean_dec(x_276); +x_282 = lean_box(0); +lean_inc(x_2); +x_283 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_279, x_265, x_281, x_282, x_9); +x_284 = lean_ctor_get(x_283, 0); +lean_inc(x_284); +x_285 = lean_ctor_get(x_283, 1); +lean_inc(x_285); +lean_dec(x_283); +x_286 = lean_ctor_get(x_284, 0); +lean_inc(x_286); +lean_dec(x_284); +x_287 = 1; +x_288 = lean_usize_add(x_7, x_287); +x_7 = x_288; +x_8 = x_286; +x_9 = x_285; +goto _start; } else { -uint8_t x_65; -x_65 = lean_nat_dec_le(x_8, x_8); -if (x_65 == 0) +lean_object* x_290; uint8_t x_291; +x_290 = lean_ctor_get(x_280, 0); +lean_inc(x_290); +lean_dec(x_280); +x_291 = l___private_Lean_Environment_0__Lean_equivInfo(x_290, x_276); +lean_dec(x_276); +lean_dec(x_290); +if (x_291 == 0) { -lean_dec(x_8); -x_14 = x_9; -goto block_64; +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +lean_dec(x_281); +lean_dec(x_279); +x_292 = l_Lean_throwAlreadyImported___rarg(x_1, x_265, x_2, x_12, x_9); +lean_dec(x_2); +x_293 = lean_ctor_get(x_292, 0); +lean_inc(x_293); +x_294 = lean_ctor_get(x_292, 1); +lean_inc(x_294); +if (lean_is_exclusive(x_292)) { + lean_ctor_release(x_292, 0); + lean_ctor_release(x_292, 1); + x_295 = x_292; +} else { + lean_dec_ref(x_292); + x_295 = lean_box(0); +} +if (lean_is_scalar(x_295)) { + x_296 = lean_alloc_ctor(1, 2, 0); +} else { + x_296 = x_295; +} +lean_ctor_set(x_296, 0, x_293); +lean_ctor_set(x_296, 1, x_294); +return x_296; } else { -size_t x_66; size_t x_67; lean_object* x_68; -x_66 = 0; -x_67 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_68 = l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8(x_7, x_66, x_67, x_9); -x_14 = x_68; -goto block_64; +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; size_t x_302; size_t x_303; +x_297 = lean_box(0); +lean_inc(x_2); +x_298 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_12, x_2, x_279, x_265, x_281, x_297, x_9); +x_299 = lean_ctor_get(x_298, 0); +lean_inc(x_299); +x_300 = lean_ctor_get(x_298, 1); +lean_inc(x_300); +lean_dec(x_298); +x_301 = lean_ctor_get(x_299, 0); +lean_inc(x_301); +lean_dec(x_299); +x_302 = 1; +x_303 = lean_usize_add(x_7, x_302); +x_7 = x_303; +x_8 = x_301; +x_9 = x_300; +goto _start; } } -block_64: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_15 = lean_unsigned_to_nat(4u); -x_16 = lean_nat_mul(x_14, x_15); -lean_dec(x_14); -x_17 = lean_unsigned_to_nat(3u); -x_18 = lean_nat_div(x_16, x_17); -lean_dec(x_16); -x_19 = l_Nat_nextPowerOfTwo_go(x_18, x_12, lean_box(0)); -lean_dec(x_18); -x_20 = lean_mk_array(x_19, x_11); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_9); -lean_ctor_set(x_21, 1, x_20); -lean_inc(x_21); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___spec__7(x_1, x_7, x_13, x_13, x_22, x_9, lean_box(0), lean_box(0), x_6); -lean_dec(x_13); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = 0; -x_29 = l_Lean_mkEmptyEnvironment___lambda__1___closed__5; -x_30 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_30, 0, x_27); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set_uint8(x_30, sizeof(void*)*2, x_28); -x_31 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_25); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = l_Array_isEmpty___rarg(x_2); -x_35 = lean_ctor_get(x_1, 3); -lean_inc(x_35); -x_36 = lean_ctor_get(x_1, 1); -lean_inc(x_36); -lean_dec(x_1); -if (x_34 == 0) -{ -uint8_t x_55; -x_55 = 1; -x_37 = x_55; -goto block_54; } -else -{ -x_37 = x_28; -goto block_54; } -block_54: -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_38 = lean_box(0); -lean_inc(x_7); -x_39 = lean_alloc_ctor(0, 5, 5); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_2); -lean_ctor_set(x_39, 2, x_35); -lean_ctor_set(x_39, 3, x_36); -lean_ctor_set(x_39, 4, x_7); -lean_ctor_set_uint32(x_39, sizeof(void*)*5, x_4); -lean_ctor_set_uint8(x_39, sizeof(void*)*5 + 4, x_37); -x_40 = l_Lean_NameSet_empty; -x_41 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_41, 0, x_26); -lean_ctor_set(x_41, 1, x_30); -lean_ctor_set(x_41, 2, x_32); -lean_ctor_set(x_41, 3, x_40); -lean_ctor_set(x_41, 4, x_39); -lean_inc(x_7); -x_42 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_41, x_7, x_9, x_33); -if (x_5 == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = lean_box(0); -x_46 = l_Lean_finalizeImport___lambda__2(x_7, x_3, x_5, x_43, x_45, x_44); -return x_46; } -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_47 = lean_ctor_get(x_42, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_42, 1); -lean_inc(x_48); -lean_dec(x_42); -x_49 = lean_runtime_mark_persistent(x_47, x_48); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_box(0); -x_53 = l_Lean_finalizeImport___lambda__2(x_7, x_3, x_5, x_50, x_52, x_51); -return x_53; } } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8) { +_start: { -uint8_t x_56; -lean_dec(x_30); -lean_dec(x_26); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_56 = !lean_is_exclusive(x_31); -if (x_56 == 0) +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_6, x_5); +if (x_9 == 0) { -return x_31; +lean_object* x_10; +lean_dec(x_1); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_8); +return x_10; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_31, 0); -x_58 = lean_ctor_get(x_31, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_31); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; -} -} -} -else +lean_object* x_11; uint8_t x_12; +x_11 = lean_array_uget(x_4, x_6); +x_12 = !lean_is_exclusive(x_7); +if (x_12 == 0) { -uint8_t x_60; -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_60 = !lean_is_exclusive(x_23); -if (x_60 == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; +x_13 = lean_ctor_get(x_7, 0); +x_14 = lean_ctor_get(x_7, 1); +x_15 = lean_array_get_size(x_14); +x_16 = l_Lean_Name_hash___override(x_11); +x_17 = 32; +x_18 = lean_uint64_shift_right(x_16, x_17); +x_19 = lean_uint64_xor(x_16, x_18); +x_20 = 16; +x_21 = lean_uint64_shift_right(x_19, x_20); +x_22 = lean_uint64_xor(x_19, x_21); +x_23 = lean_uint64_to_usize(x_22); +x_24 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_25 = 1; +x_26 = lean_usize_sub(x_24, x_25); +x_27 = lean_usize_land(x_23, x_26); +x_28 = lean_array_uget(x_14, x_27); +x_29 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_11, x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_13, x_30); +lean_dec(x_13); +lean_inc(x_1); +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_11); +lean_ctor_set(x_32, 1, x_1); +lean_ctor_set(x_32, 2, x_28); +x_33 = lean_array_uset(x_14, x_27, x_32); +x_34 = lean_unsigned_to_nat(4u); +x_35 = lean_nat_mul(x_31, x_34); +x_36 = lean_unsigned_to_nat(3u); +x_37 = lean_nat_div(x_35, x_36); +lean_dec(x_35); +x_38 = lean_array_get_size(x_33); +x_39 = lean_nat_dec_le(x_37, x_38); +lean_dec(x_38); +lean_dec(x_37); +if (x_39 == 0) { -return x_23; +lean_object* x_40; size_t x_41; +x_40 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(x_33); +lean_ctor_set(x_7, 1, x_40); +lean_ctor_set(x_7, 0, x_31); +x_41 = lean_usize_add(x_6, x_25); +x_6 = x_41; +goto _start; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_23, 0); -x_62 = lean_ctor_get(x_23, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_23); -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_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +size_t x_43; +lean_ctor_set(x_7, 1, x_33); +lean_ctor_set(x_7, 0, x_31); +x_43 = lean_usize_add(x_6, x_25); +x_6 = x_43; +goto _start; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___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) { -_start: +else { -lean_object* x_8; -x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -return x_8; +size_t x_45; +lean_dec(x_28); +lean_dec(x_11); +x_45 = lean_usize_add(x_6, x_25); +x_6 = x_45; +goto _start; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___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) { -_start: +else { -size_t x_10; size_t x_11; lean_object* x_12; -x_10 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_11 = lean_unbox_usize(x_7); +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; uint64_t x_56; size_t x_57; size_t x_58; size_t x_59; size_t x_60; size_t x_61; lean_object* x_62; uint8_t x_63; +x_47 = lean_ctor_get(x_7, 0); +x_48 = lean_ctor_get(x_7, 1); +lean_inc(x_48); +lean_inc(x_47); lean_dec(x_7); -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_2, x_3, x_4, x_5, x_10, x_11, x_8, x_9); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___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) { -_start: +x_49 = lean_array_get_size(x_48); +x_50 = l_Lean_Name_hash___override(x_11); +x_51 = 32; +x_52 = lean_uint64_shift_right(x_50, x_51); +x_53 = lean_uint64_xor(x_50, x_52); +x_54 = 16; +x_55 = lean_uint64_shift_right(x_53, x_54); +x_56 = lean_uint64_xor(x_53, x_55); +x_57 = lean_uint64_to_usize(x_56); +x_58 = lean_usize_of_nat(x_49); +lean_dec(x_49); +x_59 = 1; +x_60 = lean_usize_sub(x_58, x_59); +x_61 = lean_usize_land(x_57, x_60); +x_62 = lean_array_uget(x_48, x_61); +x_63 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_11, x_62); +if (x_63 == 0) { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_10 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_11 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_1, x_2, x_3, x_4, x_9, x_10, x_7, x_8); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___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_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_64 = lean_unsigned_to_nat(1u); +x_65 = lean_nat_add(x_47, x_64); +lean_dec(x_47); +lean_inc(x_1); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_11); +lean_ctor_set(x_66, 1, x_1); +lean_ctor_set(x_66, 2, x_62); +x_67 = lean_array_uset(x_48, x_61, x_66); +x_68 = lean_unsigned_to_nat(4u); +x_69 = lean_nat_mul(x_65, x_68); +x_70 = lean_unsigned_to_nat(3u); +x_71 = lean_nat_div(x_69, x_70); +lean_dec(x_69); +x_72 = lean_array_get_size(x_67); +x_73 = lean_nat_dec_le(x_71, x_72); +lean_dec(x_72); +lean_dec(x_71); +if (x_73 == 0) { -lean_object* x_10; -x_10 = l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} +lean_object* x_74; lean_object* x_75; size_t x_76; +x_74 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_finalizeImport___spec__2(x_67); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_65); +lean_ctor_set(x_75, 1, x_74); +x_76 = lean_usize_add(x_6, x_59); +x_6 = x_76; +x_7 = x_75; +goto _start; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; +lean_object* x_78; size_t x_79; +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_65); +lean_ctor_set(x_78, 1, x_67); +x_79 = lean_usize_add(x_6, x_59); +x_6 = x_79; +x_7 = x_78; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_finalizeImport___lambda__1(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; -} +lean_object* x_81; size_t x_82; +lean_dec(x_62); +lean_dec(x_11); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_47); +lean_ctor_set(x_81, 1, x_48); +x_82 = lean_usize_add(x_6, x_59); +x_6 = x_82; +x_7 = x_81; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_3); -lean_dec(x_3); -x_8 = l_Lean_finalizeImport___lambda__2(x_1, x_2, x_7, x_4, x_5, x_6); -lean_dec(x_5); -return x_8; } } -LEAN_EXPORT lean_object* l_Lean_finalizeImport___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint32_t x_7; uint8_t x_8; lean_object* x_9; -x_7 = lean_unbox_uint32(x_4); -lean_dec(x_4); -x_8 = lean_unbox(x_5); -lean_dec(x_5); -x_9 = l_Lean_finalizeImport(x_1, x_2, x_3, x_7, x_8, x_6); -return x_9; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___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_1; -x_1 = lean_mk_string_unchecked("import failed, trying to import module with anonymous name", 58, 58); -return x_1; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__2() { -_start: +lean_object* x_10; uint8_t x_11; +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_nat_dec_lt(x_6, x_10); +if (x_11 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__1; -x_2 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +lean_object* x_12; +lean_dec(x_6); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_9); +return x_12; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_5, x_4); -if (x_8 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_5); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; +x_14 = lean_array_fget(x_2, x_6); +x_15 = lean_ctor_get(x_14, 2); +lean_inc(x_15); +x_16 = lean_array_get_size(x_15); +x_17 = lean_unsigned_to_nat(0u); +x_18 = l_Array_toSubarray___rarg(x_15, x_17, x_16); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_5); +x_22 = lean_array_size(x_19); +x_23 = 0; +lean_inc(x_6); +x_24 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_6, x_19, x_20, x_19, x_22, x_23, x_21, x_9); +lean_dec(x_19); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_9; -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_6); -lean_ctor_set(x_9, 1, x_7); -return x_9; -} -else +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 0) { -lean_object* x_10; lean_object* x_11; +lean_object* x_29; lean_object* x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_29 = lean_ctor_get(x_26, 0); +x_30 = lean_ctor_get(x_14, 3); +lean_inc(x_30); +lean_dec(x_14); +x_31 = lean_array_size(x_30); +lean_inc(x_6); +x_32 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_6, x_20, x_30, x_30, x_31, x_23, x_29, x_27); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_ctor_set(x_26, 0, x_33); +x_35 = lean_ctor_get(x_4, 2); +x_36 = lean_nat_add(x_6, x_35); lean_dec(x_6); -x_10 = lean_array_uget(x_3, x_5); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__2; -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_7); -return x_13; +x_5 = x_26; +x_6 = x_36; +x_7 = lean_box(0); +x_8 = lean_box(0); +x_9 = x_34; +goto _start; } else { -size_t x_14; size_t x_15; lean_object* x_16; -lean_dec(x_11); -x_14 = 1; -x_15 = lean_usize_add(x_5, x_14); -x_16 = lean_box(0); -x_5 = x_15; -x_6 = x_16; +lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_38 = lean_ctor_get(x_26, 0); +x_39 = lean_ctor_get(x_26, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_26); +x_40 = lean_ctor_get(x_14, 3); +lean_inc(x_40); +lean_dec(x_14); +x_41 = lean_array_size(x_40); +lean_inc(x_6); +x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_6, x_20, x_40, x_40, x_41, x_23, x_38, x_27); +lean_dec(x_40); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_39); +x_46 = lean_ctor_get(x_4, 2); +x_47 = lean_nat_add(x_6, x_46); +lean_dec(x_6); +x_5 = x_45; +x_6 = x_47; +x_7 = lean_box(0); +x_8 = lean_box(0); +x_9 = x_44; goto _start; } } -} -} -LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1(lean_object* x_1, lean_object* x_2, uint32_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = l_Lean_finalizeImport(x_7, x_1, x_2, x_3, x_4, x_6); -return x_8; -} -} -static lean_object* _init_l_Lean_importModules___lambda__2___closed__1() { -_start: +uint8_t x_49; +lean_dec(x_14); +lean_dec(x_6); +x_49 = !lean_is_exclusive(x_24); +if (x_49 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; -x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; -x_3 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 2, x_2); -lean_ctor_set(x_3, 3, x_2); -return x_3; -} +return x_24; } -LEAN_EXPORT lean_object* l_Lean_importModules___lambda__2(lean_object* x_1, lean_object* x_2, uint32_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -lean_inc(x_1); -x_7 = lean_alloc_closure((void*)(l_Lean_importModulesCore___boxed), 3, 1); -lean_closure_set(x_7, 0, x_1); -x_8 = l_Lean_importModules___lambda__2___closed__1; -x_9 = lean_alloc_closure((void*)(l_Lean_ImportStateM_run___rarg), 3, 2); -lean_closure_set(x_9, 0, x_7); -lean_closure_set(x_9, 1, x_8); -x_10 = lean_box_uint32(x_3); -x_11 = lean_box(x_4); -x_12 = lean_alloc_closure((void*)(l_Lean_importModules___lambda__1___boxed), 6, 4); -lean_closure_set(x_12, 0, x_1); -lean_closure_set(x_12, 1, x_2); -lean_closure_set(x_12, 2, x_10); -lean_closure_set(x_12, 3, x_11); -x_13 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_13, 0, x_9); -lean_closure_set(x_13, 1, x_12); -x_14 = l_Lean_withImporting___rarg(x_13, x_6); -return x_14; -} +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_24, 0); +x_51 = lean_ctor_get(x_24, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_24); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } -static lean_object* _init_l_Lean_importModules___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("import", 6, 6); -return x_1; } } -static lean_object* _init_l_Lean_importModules___boxed__const__1() { -_start: +else { -size_t x_1; lean_object* x_2; -x_1 = 0; -x_2 = lean_box_usize(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* lean_import_modules(lean_object* x_1, lean_object* x_2, uint32_t x_3, uint8_t x_4, lean_object* x_5) { -_start: +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; size_t x_64; size_t x_65; lean_object* x_66; +x_53 = lean_ctor_get(x_5, 0); +x_54 = lean_ctor_get(x_5, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_5); +x_55 = lean_array_fget(x_2, x_6); +x_56 = lean_ctor_get(x_55, 2); +lean_inc(x_56); +x_57 = lean_array_get_size(x_56); +x_58 = lean_unsigned_to_nat(0u); +x_59 = l_Array_toSubarray___rarg(x_56, x_58, x_57); +x_60 = lean_ctor_get(x_55, 1); +lean_inc(x_60); +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_53); +lean_ctor_set(x_62, 1, x_54); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_59); +lean_ctor_set(x_63, 1, x_62); +x_64 = lean_array_size(x_60); +x_65 = 0; +lean_inc(x_6); +x_66 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_6, x_60, x_61, x_60, x_64, x_65, x_63, x_9); +lean_dec(x_60); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_6 = lean_box(0); -x_7 = lean_array_size(x_1); -x_8 = lean_box(0); -x_9 = lean_box_usize(x_7); -x_10 = l_Lean_importModules___boxed__const__1; -lean_inc_n(x_1, 2); -x_11 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___boxed), 7, 6); -lean_closure_set(x_11, 0, x_1); -lean_closure_set(x_11, 1, x_6); -lean_closure_set(x_11, 2, x_1); -lean_closure_set(x_11, 3, x_9); -lean_closure_set(x_11, 4, x_10); -lean_closure_set(x_11, 5, x_8); -x_12 = lean_box_uint32(x_3); -x_13 = lean_box(x_4); -lean_inc(x_2); -x_14 = lean_alloc_closure((void*)(l_Lean_importModules___lambda__2___boxed), 6, 4); -lean_closure_set(x_14, 0, x_1); -lean_closure_set(x_14, 1, x_2); -lean_closure_set(x_14, 2, x_12); -lean_closure_set(x_14, 3, x_13); -x_15 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_15, 0, x_11); -lean_closure_set(x_15, 1, x_14); -x_16 = l_Lean_importModules___closed__1; -x_17 = lean_box(0); -x_18 = l_Lean_profileitIOUnsafe___rarg(x_16, x_2, x_15, x_17, x_5); -lean_dec(x_2); -return x_18; +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; 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_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +lean_dec(x_67); +x_69 = lean_ctor_get(x_66, 1); +lean_inc(x_69); +lean_dec(x_66); +x_70 = lean_ctor_get(x_68, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_68, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_72 = x_68; +} else { + lean_dec_ref(x_68); + x_72 = lean_box(0); +} +x_73 = lean_ctor_get(x_55, 3); +lean_inc(x_73); +lean_dec(x_55); +x_74 = lean_array_size(x_73); +lean_inc(x_6); +x_75 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_6, x_61, x_73, x_73, x_74, x_65, x_70, x_69); +lean_dec(x_73); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec(x_75); +if (lean_is_scalar(x_72)) { + x_78 = lean_alloc_ctor(0, 2, 0); +} else { + x_78 = x_72; } +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_71); +x_79 = lean_ctor_get(x_4, 2); +x_80 = lean_nat_add(x_6, x_79); +lean_dec(x_6); +x_5 = x_78; +x_6 = x_80; +x_7 = lean_box(0); +x_8 = lean_box(0); +x_9 = x_77; +goto _start; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_9 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1(x_1, x_2, x_3, x_8, x_9, x_6, x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_55); +lean_dec(x_6); +x_82 = lean_ctor_get(x_66, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_66, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_84 = x_66; +} else { + lean_dec_ref(x_66); + x_84 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint32_t x_7; uint8_t x_8; lean_object* x_9; -x_7 = lean_unbox_uint32(x_3); -lean_dec(x_3); -x_8 = lean_unbox(x_4); -lean_dec(x_4); -x_9 = l_Lean_importModules___lambda__1(x_1, x_2, x_7, x_8, x_5, x_6); -return x_9; +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_84; } +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; } -LEAN_EXPORT lean_object* l_Lean_importModules___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint32_t x_7; uint8_t x_8; lean_object* x_9; -x_7 = lean_unbox_uint32(x_3); -lean_dec(x_3); -x_8 = lean_unbox(x_4); -lean_dec(x_4); -x_9 = l_Lean_importModules___lambda__2(x_1, x_2, x_7, x_8, x_5, x_6); -lean_dec(x_5); -return x_9; } } -LEAN_EXPORT lean_object* l_Lean_importModules___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -uint32_t x_6; uint8_t x_7; lean_object* x_8; -x_6 = lean_unbox_uint32(x_3); -lean_dec(x_3); -x_7 = lean_unbox(x_4); -lean_dec(x_4); -x_8 = lean_import_modules(x_1, x_2, x_6, x_7, x_5); -return x_8; } } -LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg(lean_object* x_1, lean_object* x_2, uint32_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { -uint8_t x_6; lean_object* x_7; -x_6 = 0; -x_7 = lean_import_modules(x_1, x_2, x_3, x_6, x_5); -if (lean_obj_tag(x_7) == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_ctor_get(x_6, 2); +lean_inc(x_7); +x_8 = lean_array_get_size(x_7); lean_dec(x_7); -lean_inc(x_8); -x_10 = lean_apply_2(x_4, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); +x_9 = lean_nat_add(x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +x_10 = lean_ctor_get(x_6, 3); +lean_inc(x_10); +lean_dec(x_6); +x_11 = lean_array_get_size(x_10); lean_dec(x_10); -x_13 = lean_environment_free_regions(x_8, x_12); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_13, 0); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_11); -return x_13; -} -else -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_11); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -} -else -{ -uint8_t x_18; +x_12 = lean_nat_add(x_9, x_11); lean_dec(x_11); -x_18 = !lean_is_exclusive(x_13); -if (x_18 == 0) -{ -return x_13; +lean_dec(x_9); +x_13 = 1; +x_14 = lean_usize_add(x_2, x_13); +x_2 = x_14; +x_4 = x_12; +goto _start; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 0); -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_13); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +return x_4; } } } -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_10, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_10, 1); -lean_inc(x_23); -lean_dec(x_10); -x_24 = lean_environment_free_regions(x_8, x_23); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_26; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -lean_ctor_set_tag(x_24, 1); -lean_ctor_set(x_24, 0, x_22); -return x_24; +lean_object* x_4; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +return x_4; } -else +} +static lean_object* _init_l_Lean_finalizeImport___lambda__2___closed__1() { +_start: { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_22); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_finalizeImport___lambda__1___boxed), 3, 0); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -uint8_t x_29; -lean_dec(x_22); -x_29 = !lean_is_exclusive(x_24); -if (x_29 == 0) +lean_object* x_7; lean_object* x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop(x_1, x_2, x_7, x_4, x_6); +if (lean_obj_tag(x_8) == 0) { -return x_24; +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_finalizeImport___lambda__2___closed__1; +if (x_3 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_box(0); +x_13 = lean_apply_3(x_11, x_9, x_12, x_10); +return x_13; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_24, 0); -x_31 = lean_ctor_get(x_24, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_24); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_runtime_mark_persistent(x_9, x_10); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_box(0); +x_18 = lean_apply_3(x_11, x_15, x_17, x_16); +return x_18; } } else { -uint8_t x_33; -lean_dec(x_4); -x_33 = !lean_is_exclusive(x_7); -if (x_33 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_8); +if (x_19 == 0) { -return x_7; +return x_8; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_7, 0); -x_35 = lean_ctor_get(x_7, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_7); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } } -LEAN_EXPORT lean_object* l_Lean_withImportModules(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_finalizeImport(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint32_t x_4, uint8_t x_5, lean_object* x_6) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_withImportModules___rarg___boxed), 5, 0); -return x_2; +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_7 = lean_ctor_get(x_1, 2); +lean_inc(x_7); +x_8 = lean_array_get_size(x_7); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_lt(x_9, x_8); +x_11 = lean_box(0); +x_12 = lean_unsigned_to_nat(1u); +lean_inc(x_8); +x_13 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_13, 0, x_9); +lean_ctor_set(x_13, 1, x_8); +lean_ctor_set(x_13, 2, x_12); +if (x_10 == 0) +{ +lean_dec(x_8); +x_14 = x_9; +goto block_65; } +else +{ +uint8_t x_66; +x_66 = lean_nat_dec_le(x_8, x_8); +if (x_66 == 0) +{ +lean_dec(x_8); +x_14 = x_9; +goto block_65; } -LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -uint32_t x_6; lean_object* x_7; -x_6 = lean_unbox_uint32(x_3); -lean_dec(x_3); -x_7 = l_Lean_withImportModules___rarg(x_1, x_2, x_6, x_4, x_5); -return x_7; +size_t x_67; size_t x_68; lean_object* x_69; +x_67 = 0; +x_68 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_69 = l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8(x_7, x_67, x_68, x_9); +x_14 = x_69; +goto block_65; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { -_start: -{ -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; uint8_t x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_10 = !lean_is_exclusive(x_5); -if (x_10 == 0) +block_65: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; size_t x_21; size_t x_22; size_t x_23; size_t x_24; lean_object* x_25; uint8_t x_26; -x_11 = lean_ctor_get(x_5, 0); -x_12 = lean_ctor_get(x_5, 1); -x_13 = lean_array_get_size(x_12); -x_14 = l_Lean_Name_hash___override(x_7); -x_15 = 32; -x_16 = lean_uint64_shift_right(x_14, x_15); -x_17 = lean_uint64_xor(x_14, x_16); -x_18 = 16; -x_19 = lean_uint64_shift_right(x_17, x_18); -x_20 = lean_uint64_xor(x_17, x_19); -x_21 = lean_uint64_to_usize(x_20); -x_22 = lean_usize_of_nat(x_13); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_15 = lean_unsigned_to_nat(4u); +x_16 = lean_nat_mul(x_14, x_15); +lean_dec(x_14); +x_17 = lean_unsigned_to_nat(3u); +x_18 = lean_nat_div(x_16, x_17); +lean_dec(x_16); +x_19 = l_Nat_nextPowerOfTwo_go(x_18, x_12, lean_box(0)); +lean_dec(x_18); +x_20 = lean_mk_array(x_19, x_11); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_9); +lean_ctor_set(x_21, 1, x_20); +lean_inc(x_21); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___spec__7(x_1, x_7, x_13, x_13, x_22, x_9, lean_box(0), lean_box(0), x_6); lean_dec(x_13); -x_23 = lean_usize_sub(x_22, x_8); -x_24 = lean_usize_land(x_21, x_23); -x_25 = lean_array_uget(x_12, x_24); -x_26 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(x_7, x_25); -if (x_26 == 0) +if (lean_obj_tag(x_23) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_27 = lean_unsigned_to_nat(1u); -x_28 = lean_nat_add(x_11, x_27); -lean_dec(x_11); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_30, 0, x_7); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = 0; +x_29 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +x_30 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_30, 0, x_27); lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_25); -x_31 = lean_array_uset(x_12, x_24, x_30); -x_32 = lean_unsigned_to_nat(4u); -x_33 = lean_nat_mul(x_28, x_32); -x_34 = lean_unsigned_to_nat(3u); -x_35 = lean_nat_div(x_33, x_34); -lean_dec(x_33); -x_36 = lean_array_get_size(x_31); -x_37 = lean_nat_dec_le(x_35, x_36); -lean_dec(x_36); -lean_dec(x_35); -if (x_37 == 0) +lean_ctor_set_uint8(x_30, sizeof(void*)*2, x_28); +x_31 = l_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates(x_25); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_38; -x_38 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(x_31); -lean_ctor_set(x_5, 1, x_38); -lean_ctor_set(x_5, 0, x_28); -x_3 = x_9; -goto _start; -} -else +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; uint8_t x_39; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Array_isEmpty___rarg(x_2); +x_35 = lean_ctor_get(x_1, 3); +lean_inc(x_35); +x_36 = lean_ctor_get(x_1, 1); +lean_inc(x_36); +lean_dec(x_1); +x_37 = lean_box(0); +lean_inc(x_7); +x_38 = lean_alloc_ctor(0, 5, 4); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_2); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 4, x_7); +lean_ctor_set_uint32(x_38, sizeof(void*)*5, x_4); +if (x_34 == 0) { -lean_ctor_set(x_5, 1, x_31); -lean_ctor_set(x_5, 0, x_28); -x_3 = x_9; -goto _start; -} +uint8_t x_56; +x_56 = 1; +x_39 = x_56; +goto block_55; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_inc(x_1); -x_41 = lean_array_uset(x_12, x_24, x_1); -x_42 = lean_box(0); -x_43 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(x_7, x_42, x_25); -x_44 = lean_array_uset(x_41, x_24, x_43); -lean_ctor_set(x_5, 1, x_44); -x_3 = x_9; -goto _start; -} +x_39 = x_28; +goto block_55; } -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; uint8_t x_61; -x_46 = lean_ctor_get(x_5, 0); -x_47 = lean_ctor_get(x_5, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_5); -x_48 = lean_array_get_size(x_47); -x_49 = l_Lean_Name_hash___override(x_7); -x_50 = 32; -x_51 = lean_uint64_shift_right(x_49, x_50); -x_52 = lean_uint64_xor(x_49, x_51); -x_53 = 16; -x_54 = lean_uint64_shift_right(x_52, x_53); -x_55 = lean_uint64_xor(x_52, x_54); -x_56 = lean_uint64_to_usize(x_55); -x_57 = lean_usize_of_nat(x_48); -lean_dec(x_48); -x_58 = lean_usize_sub(x_57, x_8); -x_59 = lean_usize_land(x_56, x_58); -x_60 = lean_array_uget(x_47, x_59); -x_61 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(x_7, x_60); -if (x_61 == 0) +block_55: { -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_62 = lean_unsigned_to_nat(1u); -x_63 = lean_nat_add(x_46, x_62); -lean_dec(x_46); -x_64 = lean_box(0); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_7); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_60); -x_66 = lean_array_uset(x_47, x_59, x_65); -x_67 = lean_unsigned_to_nat(4u); -x_68 = lean_nat_mul(x_63, x_67); -x_69 = lean_unsigned_to_nat(3u); -x_70 = lean_nat_div(x_68, x_69); -lean_dec(x_68); -x_71 = lean_array_get_size(x_66); -x_72 = lean_nat_dec_le(x_70, x_71); -lean_dec(x_71); -lean_dec(x_70); -if (x_72 == 0) +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = l_Lean_Kernel_instInhabitedDiagnostics___closed__3; +x_41 = l_Lean_NameSet_empty; +x_42 = lean_alloc_ctor(0, 6, 1); +lean_ctor_set(x_42, 0, x_30); +lean_ctor_set(x_42, 1, x_40); +lean_ctor_set(x_42, 2, x_26); +lean_ctor_set(x_42, 3, x_32); +lean_ctor_set(x_42, 4, x_41); +lean_ctor_set(x_42, 5, x_38); +lean_ctor_set_uint8(x_42, sizeof(void*)*6, x_39); +lean_inc(x_7); +x_43 = l___private_Lean_Environment_0__Lean_setImportedEntries(x_42, x_7, x_9, x_33); +if (x_5 == 0) { -lean_object* x_73; lean_object* x_74; -x_73 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(x_66); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_63); -lean_ctor_set(x_74, 1, x_73); -x_3 = x_9; -x_5 = x_74; -goto _start; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_box(0); +x_47 = l_Lean_finalizeImport___lambda__2(x_7, x_3, x_5, x_44, x_46, x_45); +return x_47; } else { -lean_object* x_76; -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_63); -lean_ctor_set(x_76, 1, x_66); -x_3 = x_9; -x_5 = x_76; -goto _start; +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_48 = lean_ctor_get(x_43, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_43, 1); +lean_inc(x_49); +lean_dec(x_43); +x_50 = lean_runtime_mark_persistent(x_48, x_49); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_box(0); +x_54 = l_Lean_finalizeImport___lambda__2(x_7, x_3, x_5, x_51, x_53, x_52); +return x_54; +} } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -lean_inc(x_1); -x_78 = lean_array_uset(x_47, x_59, x_1); -x_79 = lean_box(0); -x_80 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(x_7, x_79, x_60); -x_81 = lean_array_uset(x_78, x_59, x_80); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_46); -lean_ctor_set(x_82, 1, x_81); -x_3 = x_9; -x_5 = x_82; -goto _start; +uint8_t x_57; +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_31); +if (x_57 == 0) +{ +return x_31; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_31, 0); +x_59 = lean_ctor_get(x_31, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_31); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } } else { +uint8_t x_61; +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -return x_5; +x_61 = !lean_is_exclusive(x_23); +if (x_61 == 0) +{ +return x_23; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_23, 0); +x_63 = lean_ctor_get(x_23, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_23); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__3(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; size_t x_11; size_t x_12; -x_7 = lean_array_uget(x_2, x_3); -x_8 = lean_array_get_size(x_7); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_nat_dec_lt(x_9, x_8); -x_11 = 1; -x_12 = lean_usize_add(x_3, x_11); -if (x_10 == 0) -{ -lean_dec(x_8); -lean_dec(x_7); -x_3 = x_12; -goto _start; +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_finalizeImport___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___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) { +_start: { -uint8_t x_14; -x_14 = lean_nat_dec_le(x_8, x_8); -if (x_14 == 0) +lean_object* x_8; +x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___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) { +_start: { -lean_dec(x_8); +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_11 = lean_unbox_usize(x_7); lean_dec(x_7); -x_3 = x_12; -goto _start; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__5(x_1, x_2, x_3, x_4, x_5, x_10, x_11, x_8, x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___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) { +_start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = 0; -x_17 = lean_usize_of_nat(x_8); -lean_dec(x_8); -lean_inc(x_1); -x_18 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__2(x_1, x_7, x_16, x_17, x_5); -lean_dec(x_7); -x_3 = x_12; -x_5 = x_18; -goto _start; +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_10 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_11 = l_Array_forIn_x27Unsafe_loop___at_Lean_finalizeImport___spec__6(x_1, x_2, x_3, x_4, x_9, x_10, x_7, x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} } +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___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_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__8(x_1, x_5, x_6, x_4); lean_dec(x_1); -return x_5; +return x_7; +} } +LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_finalizeImport___lambda__1(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_array_get_size(x_3); -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_nat_dec_lt(x_5, x_4); -if (x_6 == 0) +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_finalizeImport___lambda__2(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_finalizeImport___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { +uint32_t x_7; uint8_t x_8; lean_object* x_9; +x_7 = lean_unbox_uint32(x_4); lean_dec(x_4); -lean_dec(x_1); -return x_2; +x_8 = lean_unbox(x_5); +lean_dec(x_5); +x_9 = l_Lean_finalizeImport(x_1, x_2, x_3, x_7, x_8, x_6); +return x_9; } -else +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__1() { +_start: { -uint8_t x_7; -x_7 = lean_nat_dec_le(x_4, x_4); -if (x_7 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("import failed, trying to import module with anonymous name", 58, 58); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__2() { +_start: { -lean_dec(x_4); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__1; +x_2 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_2, 0, x_1); return x_2; } +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_5, x_4); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_6); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} else { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = 0; -x_9 = lean_usize_of_nat(x_4); -lean_dec(x_4); -x_10 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__3(x_1, x_3, x_8, x_9, x_2); -return x_10; +lean_object* x_10; lean_object* x_11; +lean_dec(x_6); +x_10 = lean_array_uget(x_3, x_5); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___closed__2; +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_7); +return x_13; +} +else +{ +size_t x_14; size_t x_15; lean_object* x_16; +lean_dec(x_11); +x_14 = 1; +x_15 = lean_usize_add(x_5, x_14); +x_16 = lean_box(0); +x_5 = x_15; +x_6 = x_16; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1(lean_object* x_1, lean_object* x_2, uint32_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_2; -x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -if (x_2 == 0) -{ -return x_1; +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = l_Lean_finalizeImport(x_7, x_1, x_2, x_3, x_4, x_6); +return x_8; } -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) -{ -uint8_t x_4; -x_4 = 0; -lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_4); -return x_1; } -else +static lean_object* _init_l_Lean_importModules___lambda__2___closed__1() { +_start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -lean_inc(x_5); -lean_dec(x_1); -x_7 = 0; -x_8 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_8, 0, x_5); -lean_ctor_set(x_8, 1, x_6); -lean_ctor_set_uint8(x_8, sizeof(void*)*2, x_7); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_mkEmptyEnvironment___lambda__1___closed__3; +x_2 = l_Lean_instInhabitedEnvExtensionInterface___closed__1; +x_3 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +lean_ctor_set(x_3, 3, x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_importModules___lambda__2(lean_object* x_1, lean_object* x_2, uint32_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_inc(x_1); +x_7 = lean_alloc_closure((void*)(l_Lean_importModulesCore___boxed), 3, 1); +lean_closure_set(x_7, 0, x_1); +x_8 = l_Lean_importModules___lambda__2___closed__1; +x_9 = lean_alloc_closure((void*)(l_Lean_ImportStateM_run___rarg), 3, 2); +lean_closure_set(x_9, 0, x_7); +lean_closure_set(x_9, 1, x_8); +x_10 = lean_box_uint32(x_3); +x_11 = lean_box(x_4); +x_12 = lean_alloc_closure((void*)(l_Lean_importModules___lambda__1___boxed), 6, 4); +lean_closure_set(x_12, 0, x_1); +lean_closure_set(x_12, 1, x_2); +lean_closure_set(x_12, 2, x_10); +lean_closure_set(x_12, 3, x_11); +x_13 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); +lean_closure_set(x_13, 0, x_9); +lean_closure_set(x_13, 1, x_12); +x_14 = l_Lean_withImporting___rarg(x_13, x_6); +return x_14; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__6(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_importModules___closed__1() { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; uint8_t x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_10 = !lean_is_exclusive(x_5); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; size_t x_21; size_t x_22; size_t x_23; size_t x_24; lean_object* x_25; uint8_t x_26; -x_11 = lean_ctor_get(x_5, 0); -x_12 = lean_ctor_get(x_5, 1); -x_13 = lean_array_get_size(x_12); -x_14 = l_Lean_Name_hash___override(x_7); -x_15 = 32; -x_16 = lean_uint64_shift_right(x_14, x_15); -x_17 = lean_uint64_xor(x_14, x_16); -x_18 = 16; -x_19 = lean_uint64_shift_right(x_17, x_18); -x_20 = lean_uint64_xor(x_17, x_19); -x_21 = lean_uint64_to_usize(x_20); -x_22 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_23 = lean_usize_sub(x_22, x_8); -x_24 = lean_usize_land(x_21, x_23); -x_25 = lean_array_uget(x_12, x_24); -x_26 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(x_7, x_25); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_27 = lean_unsigned_to_nat(1u); -x_28 = lean_nat_add(x_11, x_27); -lean_dec(x_11); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_30, 0, x_7); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_25); -x_31 = lean_array_uset(x_12, x_24, x_30); -x_32 = lean_unsigned_to_nat(4u); -x_33 = lean_nat_mul(x_28, x_32); -x_34 = lean_unsigned_to_nat(3u); -x_35 = lean_nat_div(x_33, x_34); -lean_dec(x_33); -x_36 = lean_array_get_size(x_31); -x_37 = lean_nat_dec_le(x_35, x_36); -lean_dec(x_36); -lean_dec(x_35); -if (x_37 == 0) -{ -lean_object* x_38; -x_38 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(x_31); -lean_ctor_set(x_5, 1, x_38); -lean_ctor_set(x_5, 0, x_28); -x_3 = x_9; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("import", 6, 6); +return x_1; } -else +} +static lean_object* _init_l_Lean_importModules___boxed__const__1() { +_start: { -lean_ctor_set(x_5, 1, x_31); -lean_ctor_set(x_5, 0, x_28); -x_3 = x_9; -goto _start; +size_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_box_usize(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* lean_import_modules(lean_object* x_1, lean_object* x_2, uint32_t x_3, uint8_t x_4, lean_object* x_5) { +_start: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_inc(x_1); -x_41 = lean_array_uset(x_12, x_24, x_1); -x_42 = lean_box(0); -x_43 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(x_7, x_42, x_25); -x_44 = lean_array_uset(x_41, x_24, x_43); -lean_ctor_set(x_5, 1, x_44); -x_3 = x_9; -goto _start; +lean_object* x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_6 = lean_box(0); +x_7 = lean_array_size(x_1); +x_8 = lean_box(0); +x_9 = lean_box_usize(x_7); +x_10 = l_Lean_importModules___boxed__const__1; +lean_inc_n(x_1, 2); +x_11 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___boxed), 7, 6); +lean_closure_set(x_11, 0, x_1); +lean_closure_set(x_11, 1, x_6); +lean_closure_set(x_11, 2, x_1); +lean_closure_set(x_11, 3, x_9); +lean_closure_set(x_11, 4, x_10); +lean_closure_set(x_11, 5, x_8); +x_12 = lean_box_uint32(x_3); +x_13 = lean_box(x_4); +lean_inc(x_2); +x_14 = lean_alloc_closure((void*)(l_Lean_importModules___lambda__2___boxed), 6, 4); +lean_closure_set(x_14, 0, x_1); +lean_closure_set(x_14, 1, x_2); +lean_closure_set(x_14, 2, x_12); +lean_closure_set(x_14, 3, x_13); +x_15 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); +lean_closure_set(x_15, 0, x_11); +lean_closure_set(x_15, 1, x_14); +x_16 = l_Lean_importModules___closed__1; +x_17 = lean_box(0); +x_18 = l_Lean_profileitIOUnsafe___rarg(x_16, x_2, x_15, x_17, x_5); +lean_dec(x_2); +return x_18; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; uint8_t x_61; -x_46 = lean_ctor_get(x_5, 0); -x_47 = lean_ctor_get(x_5, 1); -lean_inc(x_47); -lean_inc(x_46); +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_9 = lean_unbox_usize(x_5); lean_dec(x_5); -x_48 = lean_array_get_size(x_47); -x_49 = l_Lean_Name_hash___override(x_7); -x_50 = 32; -x_51 = lean_uint64_shift_right(x_49, x_50); -x_52 = lean_uint64_xor(x_49, x_51); -x_53 = 16; -x_54 = lean_uint64_shift_right(x_52, x_53); -x_55 = lean_uint64_xor(x_52, x_54); -x_56 = lean_uint64_to_usize(x_55); -x_57 = lean_usize_of_nat(x_48); -lean_dec(x_48); -x_58 = lean_usize_sub(x_57, x_8); -x_59 = lean_usize_land(x_56, x_58); -x_60 = lean_array_uget(x_47, x_59); -x_61 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(x_7, x_60); -if (x_61 == 0) -{ -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_62 = lean_unsigned_to_nat(1u); -x_63 = lean_nat_add(x_46, x_62); -lean_dec(x_46); -x_64 = lean_box(0); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_7); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_60); -x_66 = lean_array_uset(x_47, x_59, x_65); -x_67 = lean_unsigned_to_nat(4u); -x_68 = lean_nat_mul(x_63, x_67); -x_69 = lean_unsigned_to_nat(3u); -x_70 = lean_nat_div(x_68, x_69); -lean_dec(x_68); -x_71 = lean_array_get_size(x_66); -x_72 = lean_nat_dec_le(x_70, x_71); -lean_dec(x_71); -lean_dec(x_70); -if (x_72 == 0) -{ -lean_object* x_73; lean_object* x_74; -x_73 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(x_66); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_63); -lean_ctor_set(x_74, 1, x_73); -x_3 = x_9; -x_5 = x_74; -goto _start; +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_importModules___spec__1(x_1, x_2, x_3, x_8, x_9, x_6, x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; } -else +} +LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_76; -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_63); -lean_ctor_set(x_76, 1, x_66); -x_3 = x_9; -x_5 = x_76; -goto _start; +uint32_t x_7; uint8_t x_8; lean_object* x_9; +x_7 = lean_unbox_uint32(x_3); +lean_dec(x_3); +x_8 = lean_unbox(x_4); +lean_dec(x_4); +x_9 = l_Lean_importModules___lambda__1(x_1, x_2, x_7, x_8, x_5, x_6); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Lean_importModules___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -lean_inc(x_1); -x_78 = lean_array_uset(x_47, x_59, x_1); -x_79 = lean_box(0); -x_80 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(x_7, x_79, x_60); -x_81 = lean_array_uset(x_78, x_59, x_80); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_46); -lean_ctor_set(x_82, 1, x_81); -x_3 = x_9; -x_5 = x_82; -goto _start; -} +uint32_t x_7; uint8_t x_8; lean_object* x_9; +x_7 = lean_unbox_uint32(x_3); +lean_dec(x_3); +x_8 = lean_unbox(x_4); +lean_dec(x_4); +x_9 = l_Lean_importModules___lambda__2(x_1, x_2, x_7, x_8, x_5, x_6); +lean_dec(x_5); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Lean_importModules___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_dec(x_1); -return x_5; -} +uint32_t x_6; uint8_t x_7; lean_object* x_8; +x_6 = lean_unbox_uint32(x_3); +lean_dec(x_3); +x_7 = lean_unbox(x_4); +lean_dec(x_4); +x_8 = lean_import_modules(x_1, x_2, x_6, x_7, x_5); +return x_8; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__7(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg(lean_object* x_1, lean_object* x_2, uint32_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; size_t x_11; size_t x_12; -x_7 = lean_array_uget(x_2, x_3); -x_8 = lean_array_get_size(x_7); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_nat_dec_lt(x_9, x_8); -x_11 = 1; -x_12 = lean_usize_add(x_3, x_11); -if (x_10 == 0) +uint8_t x_6; lean_object* x_7; +x_6 = 0; +x_7 = lean_import_modules(x_1, x_2, x_3, x_6, x_5); +if (lean_obj_tag(x_7) == 0) { -lean_dec(x_8); +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); lean_dec(x_7); -x_3 = x_12; -goto _start; -} -else +lean_inc(x_8); +x_10 = lean_apply_2(x_4, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_environment_free_regions(x_8, x_12); +if (lean_obj_tag(x_13) == 0) { uint8_t x_14; -x_14 = lean_nat_dec_le(x_8, x_8); +x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { -lean_dec(x_8); -lean_dec(x_7); -x_3 = x_12; -goto _start; +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_11); +return x_13; } else { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = 0; -x_17 = lean_usize_of_nat(x_8); -lean_dec(x_8); -lean_inc(x_1); -x_18 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__6(x_1, x_7, x_16, x_17, x_5); -lean_dec(x_7); -x_3 = x_12; -x_5 = x_18; -goto _start; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } +else +{ +uint8_t x_18; +lean_dec(x_11); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) +{ +return x_13; } else { -lean_dec(x_1); -return x_5; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 0); +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_13); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_array_get_size(x_3); -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_nat_dec_lt(x_5, x_4); -if (x_6 == 0) +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_10, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_10, 1); +lean_inc(x_23); +lean_dec(x_10); +x_24 = lean_environment_free_regions(x_8, x_23); +if (lean_obj_tag(x_24) == 0) { -lean_dec(x_4); -lean_dec(x_1); -return x_2; +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +lean_ctor_set_tag(x_24, 1); +lean_ctor_set(x_24, 0, x_22); +return x_24; } else { -uint8_t x_7; -x_7 = lean_nat_dec_le(x_4, x_4); -if (x_7 == 0) +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +else { -lean_dec(x_4); -lean_dec(x_1); -return x_2; +uint8_t x_29; +lean_dec(x_22); +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) +{ +return x_24; } else { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = 0; -x_9 = lean_usize_of_nat(x_4); -lean_dec(x_4); -x_10 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__7(x_1, x_3, x_8, x_9, x_2); -return x_10; +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_24, 0); +x_31 = lean_ctor_get(x_24, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_24); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_2, x_3); -if (x_5 == 0) +else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; -x_6 = lean_array_uget(x_1, x_2); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_add(x_4, x_7); -lean_dec(x_7); +uint8_t x_33; lean_dec(x_4); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_2 = x_10; -x_4 = x_8; -goto _start; +x_33 = !lean_is_exclusive(x_7); +if (x_33 == 0) +{ +return x_7; } else { -return x_4; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_7, 0); +x_35 = lean_ctor_get(x_7, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_7); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withImportModules(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_withImportModules___rarg___boxed), 5, 0); +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint32_t x_6; lean_object* x_7; +x_6 = lean_unbox_uint32(x_3); +lean_dec(x_3); +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_6600____spec__10(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { uint8_t x_6; @@ -15875,7 +16438,7 @@ return x_5; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__11(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__3(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { uint8_t x_6; @@ -15914,7 +16477,7 @@ x_16 = 0; x_17 = lean_usize_of_nat(x_8); lean_dec(x_8); lean_inc(x_1); -x_18 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__10(x_1, x_7, x_16, x_17, x_5); +x_18 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__2(x_1, x_7, x_16, x_17, x_5); lean_dec(x_7); x_3 = x_12; x_5 = x_18; @@ -15928,1387 +16491,1008 @@ lean_dec(x_1); return x_5; } } -} -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_array_get_size(x_3); -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_nat_dec_lt(x_5, x_4); -if (x_6 == 0) -{ -lean_dec(x_4); -lean_dec(x_1); -return x_2; -} -else -{ -uint8_t x_7; -x_7 = lean_nat_dec_le(x_4, x_4); -if (x_7 == 0) -{ -lean_dec(x_4); -lean_dec(x_1); -return x_2; -} -else -{ -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = 0; -x_9 = lean_usize_of_nat(x_4); -lean_dec(x_4); -x_10 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__11(x_1, x_3, x_8, x_9, x_2); -return x_10; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; -x_3 = lean_box(0); -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_6600____lambda__2___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_unsigned_to_nat(1u); -x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); -return x_3; -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__1; -x_3 = lean_mk_array(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__2; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(0u); -x_4 = lean_nat_dec_lt(x_3, x_2); -x_5 = lean_box(0); -if (x_4 == 0) -{ -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_dec(x_2); -x_6 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__3; -x_7 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__1(x_5, x_6, x_1); -x_8 = 1; -x_9 = l_Lean_mkEmptyEnvironment___lambda__1___closed__5; -x_10 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_10, 0, x_7); -lean_ctor_set(x_10, 1, x_9); -lean_ctor_set_uint8(x_10, sizeof(void*)*2, x_8); -x_11 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__4(x_10); -return x_11; -} -else -{ -uint8_t x_12; -x_12 = lean_nat_dec_le(x_2, x_2); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_2); -x_13 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__3; -x_14 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__5(x_5, x_13, x_1); -x_15 = 1; -x_16 = l_Lean_mkEmptyEnvironment___lambda__1___closed__5; -x_17 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_17, 0, x_14); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set_uint8(x_17, sizeof(void*)*2, x_15); -x_18 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__4(x_17); -return x_18; -} -else -{ -size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_19 = 0; -x_20 = lean_usize_of_nat(x_2); -lean_dec(x_2); -x_21 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__8(x_1, x_19, x_20, x_3); -x_22 = lean_unsigned_to_nat(4u); -x_23 = lean_nat_mul(x_21, x_22); -lean_dec(x_21); -x_24 = lean_unsigned_to_nat(3u); -x_25 = lean_nat_div(x_23, x_24); -lean_dec(x_23); -x_26 = lean_unsigned_to_nat(1u); -x_27 = l_Nat_nextPowerOfTwo_go(x_25, x_26, lean_box(0)); -lean_dec(x_25); -x_28 = lean_mk_array(x_27, x_5); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_3); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__9(x_5, x_29, x_1); -x_31 = 1; -x_32 = l_Lean_mkEmptyEnvironment___lambda__1___closed__5; -x_33 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_33, 0, x_30); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set_uint8(x_33, sizeof(void*)*2, x_31); -x_34 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__4(x_33); -return x_34; -} -} -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("namespacesExt", 13, 13); -return x_1; -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Environment___hyg_2684____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____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_6600____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__1), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___boxed), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_List_toArray___rarg), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____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_6600____closed__2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__3; -x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__4; -x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__5; -x_5 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_2); -lean_ctor_set(x_5, 2, x_3); -lean_ctor_set(x_5, 3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____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_6600____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__2(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__3(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__1(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__6(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__7(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__5(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__8(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_6600____spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_array_get_size(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_4); +if (x_6 == 0) +{ lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__10(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; +lean_dec(x_1); +return x_2; } +else +{ +uint8_t x_7; +x_7 = lean_nat_dec_le(x_4, x_4); +if (x_7 == 0) +{ +lean_dec(x_4); +lean_dec(x_1); +return x_2; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_4); lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__11(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__3(x_1, x_3, x_8, x_9, x_2); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__4(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6600____spec__9(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; +uint8_t x_2; +x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +if (x_2 == 0) +{ +return x_1; } +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = 0; +lean_ctor_set_uint8(x_1, sizeof(void*)*2, x_4); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2(x_1); +lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_inc(x_5); lean_dec(x_1); -return x_2; +x_7 = 0; +x_8 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_6); +lean_ctor_set_uint8(x_8, sizeof(void*)*2, x_7); +return x_8; } } -static lean_object* _init_l_Lean_Kernel_instInhabitedDiagnostics___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_mkEmptyEnvironment___lambda__1___closed__5; -x_2 = 0; -x_3 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Kernel_instInhabitedDiagnostics() { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__6(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = l_Lean_Kernel_instInhabitedDiagnostics___closed__1; -return x_1; -} +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; uint8_t x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = !lean_is_exclusive(x_5); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; size_t x_21; size_t x_22; size_t x_23; size_t x_24; lean_object* x_25; uint8_t x_26; +x_11 = lean_ctor_get(x_5, 0); +x_12 = lean_ctor_get(x_5, 1); +x_13 = lean_array_get_size(x_12); +x_14 = l_Lean_Name_hash___override(x_7); +x_15 = 32; +x_16 = lean_uint64_shift_right(x_14, x_15); +x_17 = lean_uint64_xor(x_14, x_16); +x_18 = 16; +x_19 = lean_uint64_shift_right(x_17, x_18); +x_20 = lean_uint64_xor(x_17, x_19); +x_21 = lean_uint64_to_usize(x_20); +x_22 = lean_usize_of_nat(x_13); +lean_dec(x_13); +x_23 = lean_usize_sub(x_22, x_8); +x_24 = lean_usize_land(x_21, x_23); +x_25 = lean_array_uget(x_12, x_24); +x_26 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(x_7, x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_27 = lean_unsigned_to_nat(1u); +x_28 = lean_nat_add(x_11, x_27); +lean_dec(x_11); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_7); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_25); +x_31 = lean_array_uset(x_12, x_24, x_30); +x_32 = lean_unsigned_to_nat(4u); +x_33 = lean_nat_mul(x_28, x_32); +x_34 = lean_unsigned_to_nat(3u); +x_35 = lean_nat_div(x_33, x_34); +lean_dec(x_33); +x_36 = lean_array_get_size(x_31); +x_37 = lean_nat_dec_le(x_35, x_36); +lean_dec(x_36); +lean_dec(x_35); +if (x_37 == 0) +{ +lean_object* x_38; +x_38 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(x_31); +lean_ctor_set(x_5, 1, x_38); +lean_ctor_set(x_5, 0, x_28); +x_3 = x_9; +goto _start; } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6727____closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Kernel_instInhabitedDiagnostics___closed__1; -x_2 = lean_alloc_closure((void*)(l_EStateM_pure___rarg), 2, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; +lean_ctor_set(x_5, 1, x_31); +lean_ctor_set(x_5, 0, x_28); +x_3 = x_9; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6727_(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6727____closed__1; -x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_2, x_1); -return x_3; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_inc(x_1); +x_41 = lean_array_uset(x_12, x_24, x_1); +x_42 = lean_box(0); +x_43 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(x_7, x_42, x_25); +x_44 = lean_array_uset(x_41, x_24, x_43); +lean_ctor_set(x_5, 1, x_44); +x_3 = x_9; +goto _start; } } -LEAN_EXPORT uint8_t lean_kernel_diag_is_enabled(lean_object* x_1) { -_start: +else { -uint8_t x_2; -x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); -lean_dec(x_1); -return x_2; -} +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; uint8_t x_61; +x_46 = lean_ctor_get(x_5, 0); +x_47 = lean_ctor_get(x_5, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_5); +x_48 = lean_array_get_size(x_47); +x_49 = l_Lean_Name_hash___override(x_7); +x_50 = 32; +x_51 = lean_uint64_shift_right(x_49, x_50); +x_52 = lean_uint64_xor(x_49, x_51); +x_53 = 16; +x_54 = lean_uint64_shift_right(x_52, x_53); +x_55 = lean_uint64_xor(x_52, x_54); +x_56 = lean_uint64_to_usize(x_55); +x_57 = lean_usize_of_nat(x_48); +lean_dec(x_48); +x_58 = lean_usize_sub(x_57, x_8); +x_59 = lean_usize_land(x_56, x_58); +x_60 = lean_array_uget(x_47, x_59); +x_61 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(x_7, x_60); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_62 = lean_unsigned_to_nat(1u); +x_63 = lean_nat_add(x_46, x_62); +lean_dec(x_46); +x_64 = lean_box(0); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_7); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_60); +x_66 = lean_array_uset(x_47, x_59, x_65); +x_67 = lean_unsigned_to_nat(4u); +x_68 = lean_nat_mul(x_63, x_67); +x_69 = lean_unsigned_to_nat(3u); +x_70 = lean_nat_div(x_68, x_69); +lean_dec(x_68); +x_71 = lean_array_get_size(x_66); +x_72 = lean_nat_dec_le(x_70, x_71); +lean_dec(x_71); +lean_dec(x_70); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; +x_73 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(x_66); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_63); +lean_ctor_set(x_74, 1, x_73); +x_3 = x_9; +x_5 = x_74; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Kernel_Diagnostics_isEnabled___boxed(lean_object* x_1) { -_start: +else { -uint8_t x_2; lean_object* x_3; -x_2 = lean_kernel_diag_is_enabled(x_1); -x_3 = lean_box(x_2); -return x_3; +lean_object* x_76; +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_63); +lean_ctor_set(x_76, 1, x_66); +x_3 = x_9; +x_5 = x_76; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag___lambda__1(uint8_t x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) +else { -lean_ctor_set_uint8(x_2, sizeof(void*)*1, x_1); -return x_2; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_inc(x_1); +x_78 = lean_array_uset(x_47, x_59, x_1); +x_79 = lean_box(0); +x_80 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(x_7, x_79, x_60); +x_81 = lean_array_uset(x_78, x_59, x_80); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_46); +lean_ctor_set(x_82, 1, x_81); +x_3 = x_9; +x_5 = x_82; +goto _start; +} +} } else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -lean_dec(x_2); -x_5 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set_uint8(x_5, sizeof(void*)*1, x_1); +lean_dec(x_1); return x_5; } } } -static lean_object* _init_l_Lean_Kernel_enableDiag___closed__1() { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__7(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = l_Lean_diagExt; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag(lean_object* x_1, uint8_t x_2) { -_start: +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_box(x_2); -x_4 = lean_alloc_closure((void*)(l_Lean_Kernel_enableDiag___lambda__1___boxed), 2, 1); -lean_closure_set(x_4, 0, x_3); -x_5 = l_Lean_Kernel_enableDiag___closed__1; -x_6 = l_Lean_EnvExtension_modifyState___rarg(x_5, x_1, x_4); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; size_t x_11; size_t x_12; +x_7 = lean_array_uget(x_2, x_3); +x_8 = lean_array_get_size(x_7); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_lt(x_9, x_8); +x_11 = 1; +x_12 = lean_usize_add(x_3, x_11); +if (x_10 == 0) { -uint8_t x_3; lean_object* x_4; -x_3 = lean_unbox(x_1); -lean_dec(x_1); -x_4 = l_Lean_Kernel_enableDiag___lambda__1(x_3, x_2); -return x_4; -} +lean_dec(x_8); +lean_dec(x_7); +x_3 = x_12; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = lean_unbox(x_2); -lean_dec(x_2); -x_4 = l_Lean_Kernel_enableDiag(x_1, x_3); -return x_4; -} -} -LEAN_EXPORT uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object* x_1) { -_start: +uint8_t x_14; +x_14 = lean_nat_dec_le(x_8, x_8); +if (x_14 == 0) { -lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_2 = l_Lean_Kernel_instInhabitedDiagnostics; -x_3 = l_Lean_Kernel_enableDiag___closed__1; -x_4 = l_Lean_EnvExtension_getState___rarg(x_2, x_3, x_1); -x_5 = lean_ctor_get_uint8(x_4, sizeof(void*)*1); -lean_dec(x_4); -return x_5; -} +lean_dec(x_8); +lean_dec(x_7); +x_3 = x_12; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Kernel_isDiagnosticsEnabled___boxed(lean_object* x_1) { -_start: +else { -uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Kernel_isDiagnosticsEnabled(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = 0; +x_17 = lean_usize_of_nat(x_8); +lean_dec(x_8); +lean_inc(x_1); +x_18 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__6(x_1, x_7, x_16, x_17, x_5); +lean_dec(x_7); +x_3 = x_12; +x_5 = x_18; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Kernel_resetDiag___lambda__1(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = !lean_is_exclusive(x_1); -if (x_2 == 0) -{ -lean_object* x_3; lean_object* x_4; -x_3 = lean_ctor_get(x_1, 0); -lean_dec(x_3); -x_4 = l_Lean_mkEmptyEnvironment___lambda__1___closed__5; -lean_ctor_set(x_1, 0, x_4); -return x_1; } else { -uint8_t x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); lean_dec(x_1); -x_6 = l_Lean_mkEmptyEnvironment___lambda__1___closed__5; -x_7 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set_uint8(x_7, sizeof(void*)*1, x_5); -return x_7; +return x_5; } } } -static lean_object* _init_l_Lean_Kernel_resetDiag___closed__1() { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Kernel_resetDiag___lambda__1), 1, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Kernel_resetDiag(lean_object* x_1) { -_start: +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_array_get_size(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_4); +if (x_6 == 0) { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Kernel_enableDiag___closed__1; -x_3 = l_Lean_Kernel_resetDiag___closed__1; -x_4 = l_Lean_EnvExtension_modifyState___rarg(x_2, x_1, x_3); -return x_4; -} +lean_dec(x_4); +lean_dec(x_1); +return x_2; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); +uint8_t x_7; +x_7 = lean_nat_dec_le(x_4, x_4); if (x_7 == 0) { -lean_object* x_8; lean_dec(x_4); -x_8 = lean_box(0); -return x_8; +lean_dec(x_1); +return x_2; } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_name_eq(x_5, x_9); -lean_dec(x_9); -if (x_10 == 0) +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__7(x_1, x_3, x_8, x_9, x_2); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_add(x_4, x_7); +lean_dec(x_7); lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_12; +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_2 = x_10; +x_4 = x_8; goto _start; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_array_fget(x_2, x_4); -lean_dec(x_4); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; -} +return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__10(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -if (lean_obj_tag(x_1) == 0) +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +lean_object* x_7; size_t x_8; size_t x_9; uint8_t x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = !lean_is_exclusive(x_5); +if (x_10 == 0) { -lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_box(2); -x_11 = lean_array_get(x_10, x_5, x_9); -lean_dec(x_9); -lean_dec(x_5); -switch (lean_obj_tag(x_11)) { -case 0: +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; size_t x_21; size_t x_22; size_t x_23; size_t x_24; lean_object* x_25; uint8_t x_26; +x_11 = lean_ctor_get(x_5, 0); +x_12 = lean_ctor_get(x_5, 1); +x_13 = lean_array_get_size(x_12); +x_14 = l_Lean_Name_hash___override(x_7); +x_15 = 32; +x_16 = lean_uint64_shift_right(x_14, x_15); +x_17 = lean_uint64_xor(x_14, x_16); +x_18 = 16; +x_19 = lean_uint64_shift_right(x_17, x_18); +x_20 = lean_uint64_xor(x_17, x_19); +x_21 = lean_uint64_to_usize(x_20); +x_22 = lean_usize_of_nat(x_13); +lean_dec(x_13); +x_23 = lean_usize_sub(x_22, x_8); +x_24 = lean_usize_land(x_21, x_23); +x_25 = lean_array_uget(x_12, x_24); +x_26 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(x_7, x_25); +if (x_26 == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_27 = lean_unsigned_to_nat(1u); +x_28 = lean_nat_add(x_11, x_27); lean_dec(x_11); -x_14 = lean_name_eq(x_3, x_12); -lean_dec(x_12); -if (x_14 == 0) -{ -lean_object* x_15; -lean_dec(x_13); -lean_free_object(x_1); -x_15 = lean_box(0); -return x_15; -} -else +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_7); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_25); +x_31 = lean_array_uset(x_12, x_24, x_30); +x_32 = lean_unsigned_to_nat(4u); +x_33 = lean_nat_mul(x_28, x_32); +x_34 = lean_unsigned_to_nat(3u); +x_35 = lean_nat_div(x_33, x_34); +lean_dec(x_33); +x_36 = lean_array_get_size(x_31); +x_37 = lean_nat_dec_le(x_35, x_36); +lean_dec(x_36); +lean_dec(x_35); +if (x_37 == 0) { -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_13); -return x_1; -} +lean_object* x_38; +x_38 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(x_31); +lean_ctor_set(x_5, 1, x_38); +lean_ctor_set(x_5, 0, x_28); +x_3 = x_9; +goto _start; } -case 1: +else { -lean_object* x_16; size_t x_17; -lean_free_object(x_1); -x_16 = lean_ctor_get(x_11, 0); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_usize_shift_right(x_2, x_6); -x_1 = x_16; -x_2 = x_17; +lean_ctor_set(x_5, 1, x_31); +lean_ctor_set(x_5, 0, x_28); +x_3 = x_9; goto _start; } -default: -{ -lean_object* x_19; -lean_free_object(x_1); -x_19 = lean_box(0); -return x_19; } +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_inc(x_1); +x_41 = lean_array_uset(x_12, x_24, x_1); +x_42 = lean_box(0); +x_43 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(x_7, x_42, x_25); +x_44 = lean_array_uset(x_41, x_24, x_43); +lean_ctor_set(x_5, 1, x_44); +x_3 = x_9; +goto _start; } } else { -lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); -lean_dec(x_1); -x_21 = 5; -x_22 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_23 = lean_usize_land(x_2, x_22); -x_24 = lean_usize_to_nat(x_23); -x_25 = lean_box(2); -x_26 = lean_array_get(x_25, x_20, x_24); -lean_dec(x_24); -lean_dec(x_20); -switch (lean_obj_tag(x_26)) { -case 0: +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; uint8_t x_61; +x_46 = lean_ctor_get(x_5, 0); +x_47 = lean_ctor_get(x_5, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_5); +x_48 = lean_array_get_size(x_47); +x_49 = l_Lean_Name_hash___override(x_7); +x_50 = 32; +x_51 = lean_uint64_shift_right(x_49, x_50); +x_52 = lean_uint64_xor(x_49, x_51); +x_53 = 16; +x_54 = lean_uint64_shift_right(x_52, x_53); +x_55 = lean_uint64_xor(x_52, x_54); +x_56 = lean_uint64_to_usize(x_55); +x_57 = lean_usize_of_nat(x_48); +lean_dec(x_48); +x_58 = lean_usize_sub(x_57, x_8); +x_59 = lean_usize_land(x_56, x_58); +x_60 = lean_array_uget(x_47, x_59); +x_61 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(x_7, x_60); +if (x_61 == 0) { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_name_eq(x_3, x_27); -lean_dec(x_27); -if (x_29 == 0) +lean_object* x_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_62 = lean_unsigned_to_nat(1u); +x_63 = lean_nat_add(x_46, x_62); +lean_dec(x_46); +x_64 = lean_box(0); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_7); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_60); +x_66 = lean_array_uset(x_47, x_59, x_65); +x_67 = lean_unsigned_to_nat(4u); +x_68 = lean_nat_mul(x_63, x_67); +x_69 = lean_unsigned_to_nat(3u); +x_70 = lean_nat_div(x_68, x_69); +lean_dec(x_68); +x_71 = lean_array_get_size(x_66); +x_72 = lean_nat_dec_le(x_70, x_71); +lean_dec(x_71); +lean_dec(x_70); +if (x_72 == 0) { -lean_object* x_30; -lean_dec(x_28); -x_30 = lean_box(0); -return x_30; +lean_object* x_73; lean_object* x_74; +x_73 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(x_66); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_63); +lean_ctor_set(x_74, 1, x_73); +x_3 = x_9; +x_5 = x_74; +goto _start; } else { -lean_object* x_31; -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_28); -return x_31; +lean_object* x_76; +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_63); +lean_ctor_set(x_76, 1, x_66); +x_3 = x_9; +x_5 = x_76; +goto _start; } } -case 1: +else { -lean_object* x_32; size_t x_33; -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_usize_shift_right(x_2, x_21); -x_1 = x_32; -x_2 = x_33; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_inc(x_1); +x_78 = lean_array_uset(x_47, x_59, x_1); +x_79 = lean_box(0); +x_80 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_NameSSet_insert___spec__10(x_7, x_79, x_60); +x_81 = lean_array_uset(x_78, x_59, x_80); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_46); +lean_ctor_set(x_82, 1, x_81); +x_3 = x_9; +x_5 = x_82; goto _start; } -default: -{ -lean_object* x_35; -x_35 = lean_box(0); -return x_35; -} -} } } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_1, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 1); -lean_inc(x_37); lean_dec(x_1); -x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__3(x_36, x_37, lean_box(0), x_38, x_3); -lean_dec(x_37); -lean_dec(x_36); -return x_39; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Name_hash___override(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__2(x_1, x_4, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Diagnostics_recordUnfold___spec__6(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_5); -return x_6; -} -else -{ -lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Name_hash___override(x_9); -x_12 = lean_uint64_to_usize(x_11); -x_13 = 1; -x_14 = lean_usize_sub(x_1, x_13); -x_15 = 5; -x_16 = lean_usize_mul(x_15, x_14); -x_17 = lean_usize_shift_right(x_12, x_16); -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_5, x_18); -lean_dec(x_5); -x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(x_6, x_17, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_19; -x_6 = x_20; -goto _start; -} -} } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__7(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_7023____spec__11(lean_object* x_1, lean_object* x_2, size_t x_3, size_t 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_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_1, 1); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 0); -lean_dec(x_11); -x_12 = lean_array_push(x_5, x_3); -x_13 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_12); -return x_1; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_1); -x_14 = lean_array_push(x_5, x_3); -x_15 = lean_array_push(x_6, x_4); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -else +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) { -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_name_eq(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; size_t x_11; size_t x_12; +x_7 = lean_array_uget(x_2, x_3); +x_8 = lean_array_get_size(x_7); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_lt(x_9, x_8); +x_11 = 1; +x_12 = lean_usize_add(x_3, x_11); +if (x_10 == 0) { -lean_object* x_19; lean_object* x_20; -lean_dec(x_6); -lean_dec(x_5); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_2, x_19); -lean_dec(x_2); -x_2 = x_20; +lean_dec(x_8); +lean_dec(x_7); +x_3 = x_12; goto _start; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -if (x_22 == 0) +uint8_t x_14; +x_14 = lean_nat_dec_le(x_8, x_8); +if (x_14 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 1); -lean_dec(x_23); -x_24 = lean_ctor_get(x_1, 0); -lean_dec(x_24); -x_25 = lean_array_fset(x_5, x_2, x_3); -x_26 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -lean_ctor_set(x_1, 1, x_26); -lean_ctor_set(x_1, 0, x_25); -return x_1; +lean_dec(x_8); +lean_dec(x_7); +x_3 = x_12; +goto _start; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_1); -x_27 = lean_array_fset(x_5, x_2, x_3); -x_28 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = 0; +x_17 = lean_usize_of_nat(x_8); +lean_dec(x_8); +lean_inc(x_1); +x_18 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__10(x_1, x_7, x_16, x_17, x_5); +lean_dec(x_7); +x_3 = x_12; +x_5 = x_18; +goto _start; +} } } +else +{ +lean_dec(x_1); +return x_5; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_array_get_size(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_4); if (x_6 == 0) { -lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = 1; -x_9 = 5; -x_10 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_14 = lean_nat_dec_lt(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_dec(x_12); -lean_dec(x_5); lean_dec(x_4); -return x_1; +lean_dec(x_1); +return x_2; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_name_eq(x_4, x_19); -if (x_21 == 0) +uint8_t x_7; +x_7 = lean_nat_dec_le(x_4, x_4); +if (x_7 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_free_object(x_15); -x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_array_fset(x_17, x_12, x_23); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_24); -return x_1; +lean_dec(x_4); +lean_dec(x_1); +return x_2; } else { -lean_object* x_25; -lean_dec(x_20); -lean_dec(x_19); -lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_4); -x_25 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_25); -return x_1; +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__11(x_1, x_3, x_8, x_9, x_2); +return x_10; } } -else +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_name_eq(x_4, x_26); -if (x_28 == 0) +lean_object* x_3; lean_object* x_4; +x_3 = lean_box(0); +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_7023____lambda__2___closed__1() { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_array_fset(x_17, x_12, x_30); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); +return x_3; } -else +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__2() { +_start: { -lean_object* x_32; lean_object* x_33; -lean_dec(x_27); -lean_dec(x_26); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_4); -lean_ctor_set(x_32, 1, x_5); -x_33 = lean_array_fset(x_17, x_12, x_32); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_33); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__1; +x_3 = lean_mk_array(x_2, x_1); +return x_3; } } +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__2; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} } -case 1: +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2(lean_object* x_1) { +_start: { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) +lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_nat_dec_lt(x_3, x_2); +x_5 = lean_box(0); +if (x_4 == 0) { -lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -x_35 = lean_ctor_get(x_15, 0); -x_36 = lean_usize_shift_right(x_2, x_9); -x_37 = lean_usize_add(x_3, x_8); -x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(x_35, x_36, x_37, x_4, x_5); -lean_ctor_set(x_15, 0, x_38); -x_39 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); -return x_1; +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_2); +x_6 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__3; +x_7 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__1(x_5, x_6, x_1); +x_8 = 1; +x_9 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +x_10 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_9); +lean_ctor_set_uint8(x_10, sizeof(void*)*2, x_8); +x_11 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__4(x_10); +return x_11; } else { -lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_15, 0); -lean_inc(x_40); -lean_dec(x_15); -x_41 = lean_usize_shift_right(x_2, x_9); -x_42 = lean_usize_add(x_3, x_8); -x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(x_40, x_41, x_42, x_4, x_5); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_array_fset(x_17, x_12, x_44); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); -return x_1; -} +uint8_t x_12; +x_12 = lean_nat_dec_le(x_2, x_2); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_2); +x_13 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__3; +x_14 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__5(x_5, x_13, x_1); +x_15 = 1; +x_16 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +x_17 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_17, 0, x_14); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set_uint8(x_17, sizeof(void*)*2, x_15); +x_18 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__4(x_17); +return x_18; } -default: +else { -lean_object* x_46; lean_object* x_47; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_4); -lean_ctor_set(x_46, 1, x_5); -x_47 = lean_array_fset(x_17, x_12, x_46); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_47); -return x_1; +size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_19 = 0; +x_20 = lean_usize_of_nat(x_2); +lean_dec(x_2); +x_21 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__8(x_1, x_19, x_20, x_3); +x_22 = lean_unsigned_to_nat(4u); +x_23 = lean_nat_mul(x_21, x_22); +lean_dec(x_21); +x_24 = lean_unsigned_to_nat(3u); +x_25 = lean_nat_div(x_23, x_24); +lean_dec(x_23); +x_26 = lean_unsigned_to_nat(1u); +x_27 = l_Nat_nextPowerOfTwo_go(x_25, x_26, lean_box(0)); +lean_dec(x_25); +x_28 = lean_mk_array(x_27, x_5); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_3); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__9(x_5, x_29, x_1); +x_31 = 1; +x_32 = l_Lean_Kernel_instInhabitedDiagnostics___closed__2; +x_33 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_33, 0, x_30); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set_uint8(x_33, sizeof(void*)*2, x_31); +x_34 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__4(x_33); +return x_34; } } } } -else -{ -lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_1, 0); -lean_inc(x_48); -lean_dec(x_1); -x_49 = 1; -x_50 = 5; -x_51 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2; -x_52 = lean_usize_land(x_2, x_51); -x_53 = lean_usize_to_nat(x_52); -x_54 = lean_array_get_size(x_48); -x_55 = lean_nat_dec_lt(x_53, x_54); -lean_dec(x_54); -if (x_55 == 0) +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__1() { +_start: { -lean_object* x_56; -lean_dec(x_53); -lean_dec(x_5); -lean_dec(x_4); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_48); -return x_56; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("namespacesExt", 13, 13); +return x_1; } -else -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_array_fget(x_48, x_53); -x_58 = lean_box(0); -x_59 = lean_array_fset(x_48, x_53, x_58); -switch (lean_obj_tag(x_57)) { -case 0: -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_57, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_62 = x_57; -} else { - lean_dec_ref(x_57); - x_62 = lean_box(0); } -x_63 = lean_name_eq(x_4, x_60); -if (x_63 == 0) +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__2() { +_start: { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_array_fset(x_59, x_53, x_65); -lean_dec(x_53); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Lean_Environment___hyg_3100____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_61); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_68 = lean_alloc_ctor(0, 2, 0); -} else { - x_68 = x_62; } -lean_ctor_set(x_68, 0, x_4); -lean_ctor_set(x_68, 1, x_5); -x_69 = lean_array_fset(x_59, x_53, x_68); -lean_dec(x_53); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__1), 2, 0); +return x_1; } } -case 1: +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__4() { +_start: { -lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_71 = lean_ctor_get(x_57, 0); -lean_inc(x_71); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - x_72 = x_57; -} else { - lean_dec_ref(x_57); - x_72 = lean_box(0); -} -x_73 = lean_usize_shift_right(x_2, x_50); -x_74 = lean_usize_add(x_3, x_49); -x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(x_71, x_73, x_74, x_4, x_5); -if (lean_is_scalar(x_72)) { - x_76 = lean_alloc_ctor(1, 1, 0); -} else { - x_76 = x_72; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___boxed), 1, 0); +return x_1; } -lean_ctor_set(x_76, 0, x_75); -x_77 = lean_array_fset(x_59, x_53, x_76); -lean_dec(x_53); -x_78 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_78, 0, x_77); -return x_78; } -default: +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__5() { +_start: { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_4); -lean_ctor_set(x_79, 1, x_5); -x_80 = lean_array_fset(x_59, x_53, x_79); -lean_dec(x_53); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; -} +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_List_toArray___rarg), 1, 0); +return x_1; } } +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____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_7023____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__3; +x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__5; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; } } -else -{ -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) -{ -lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; -x_83 = lean_unsigned_to_nat(0u); -x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__7(x_1, x_83, x_4, x_5); -x_85 = 7; -x_86 = lean_usize_dec_le(x_85, x_3); -if (x_86 == 0) -{ -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_dec_lt(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023_(lean_object* x_1) { +_start: { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_84, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 1); -lean_inc(x_91); -lean_dec(x_84); -x_92 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Diagnostics_recordUnfold___spec__6(x_3, x_90, x_91, lean_box(0), x_83, x_92); -lean_dec(x_91); -lean_dec(x_90); -return x_93; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__6; +x_3 = l_Lean_registerSimplePersistentEnvExtension___rarg(x_2, x_1); +return x_3; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -return x_84; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__2(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -return x_84; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__3(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; } } -else -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; -x_94 = lean_ctor_get(x_1, 0); -x_95 = lean_ctor_get(x_1, 1); -lean_inc(x_95); -lean_inc(x_94); -lean_dec(x_1); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -x_97 = lean_unsigned_to_nat(0u); -x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__7(x_96, x_97, x_4, x_5); -x_99 = 7; -x_100 = lean_usize_dec_le(x_99, x_3); -if (x_100 == 0) -{ -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); -x_102 = lean_unsigned_to_nat(4u); -x_103 = lean_nat_dec_lt(x_101, x_102); -lean_dec(x_101); -if (x_103 == 0) +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_98, 1); -lean_inc(x_105); -lean_dec(x_98); -x_106 = l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Diagnostics_recordUnfold___spec__6(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; +lean_object* x_4; +x_4 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -return x_98; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__6(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -return x_98; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__7(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} } +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__5(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__8(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Name_hash___override(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(x_1, x_5, x_6, x_2, x_3); -return x_7; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__10(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; } } -LEAN_EXPORT lean_object* lean_kernel_record_unfold(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_3; -x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); -if (x_3 == 0) -{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__11(x_1, x_2, x_6, x_7, x_5); lean_dec(x_2); -return x_1; +return x_8; } -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_5, x_2); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_unsigned_to_nat(1u); -x_8 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_5, x_2, x_7); -lean_ctor_set(x_1, 0, x_8); -return x_1; } -else +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_6, 0); -lean_inc(x_9); -lean_dec(x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_9, x_10); -lean_dec(x_9); -x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_5, x_2, x_11); -lean_ctor_set(x_1, 0, x_12); -return x_1; +lean_object* x_4; +x_4 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_7023____spec__9(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; } } -else +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___boxed(lean_object* x_1) { +_start: { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); +lean_object* x_2; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2(x_1); lean_dec(x_1); -lean_inc(x_13); -x_14 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_13, x_2); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_unsigned_to_nat(1u); -x_16 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_13, x_2, x_15); -x_17 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_3); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_14, 0); -lean_inc(x_18); -lean_dec(x_14); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_18, x_19); -lean_dec(x_18); -x_21 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_13, x_2, x_20); -x_22 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set_uint8(x_22, sizeof(void*)*1, x_3); -return x_22; -} +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag(lean_object* x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Kernel_Environment_enableDiag(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__3(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); lean_dec(x_2); -lean_dec(x_1); -return x_6; +x_4 = l_Lean_Kernel_enableDiag(x_1, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object* x_1) { _start: { -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__2(x_1, x_4, x_3); -lean_dec(x_3); -return x_5; +uint8_t x_2; +x_2 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Kernel_isDiagnosticsEnabled___boxed(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_1, x_2); -lean_dec(x_2); +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Kernel_isDiagnosticsEnabled(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Diagnostics_recordUnfold___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Kernel_resetDiag(lean_object* x_1) { _start: { -size_t x_7; lean_object* x_8; -x_7 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Kernel_Diagnostics_recordUnfold___spec__6(x_7, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_3); -lean_dec(x_2); -return x_8; +lean_object* x_2; +x_2 = l_Lean_Kernel_Environment_resetDiag(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Kernel_getDiagnostics(lean_object* x_1) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(x_1, x_6, x_7, x_4, x_5); -return x_8; +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +return x_2; } } -LEAN_EXPORT lean_object* lean_kernel_get_diag(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Kernel_getDiagnostics___boxed(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Kernel_instInhabitedDiagnostics; -x_3 = l_Lean_Kernel_enableDiag___closed__1; -x_4 = l_Lean_EnvExtension_getState___rarg(x_2, x_3, x_1); +lean_object* x_2; +x_2 = l_Lean_Kernel_getDiagnostics(x_1); lean_dec(x_1); -return x_4; +return x_2; } } -LEAN_EXPORT lean_object* lean_kernel_set_diag(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Kernel_setDiagnostics(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; -x_3 = l_Lean_Kernel_enableDiag___closed__1; -x_4 = l_Lean_EnvExtension_setState___rarg(x_3, x_1, x_2); -return x_4; +lean_object* x_3; +x_3 = lean_kernel_set_diag(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Environment_registerNamespace___closed__1() { @@ -17418,7 +17602,7 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes_go(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 1) @@ -17450,42 +17634,126 @@ return x_1; } } } -LEAN_EXPORT lean_object* lean_environment_add(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; -x_3 = l_Lean_ConstantInfo_name(x_2); -if (lean_obj_tag(x_3) == 1) +if (lean_obj_tag(x_2) == 1) { -lean_object* x_4; lean_object* x_5; uint32_t x_6; uint32_t x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_3, 1); +lean_object* x_3; lean_object* x_4; uint32_t x_5; uint32_t x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_string_utf8_get(x_3, x_4); +lean_dec(x_3); +x_6 = 95; +x_7 = lean_uint32_dec_eq(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; +x_8 = l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes_go(x_1, x_2); +return x_8; +} +else +{ +lean_dec(x_2); +return x_1; +} +} +else +{ +lean_dec(x_2); +return x_1; +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); lean_inc(x_4); -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_string_utf8_get(x_4, x_5); -lean_dec(x_4); -x_7 = 95; -x_8 = lean_uint32_dec_eq(x_6, x_7); -if (x_8 == 0) +lean_dec(x_2); +x_5 = l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes(x_1, x_3); +x_1 = x_5; +x_2 = x_4; +goto _start; +} +} +} +static lean_object* _init_l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__1() { +_start: { -lean_object* x_9; lean_object* x_10; -x_9 = l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes(x_1, x_3); -x_10 = l___private_Lean_Environment_0__Lean_Environment_addAux(x_9, x_2); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("rec", 3, 3); +return x_1; +} +} +static lean_object* _init_l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; } else { -lean_object* x_11; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); lean_dec(x_3); -x_11 = l___private_Lean_Environment_0__Lean_Environment_addAux(x_1, x_2); -return x_11; +x_6 = l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__2; +x_7 = l_Lean_Name_append(x_5, x_6); +x_8 = l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes(x_1, x_7); +x_1 = x_8; +x_2 = x_4; +goto _start; } } +} +LEAN_EXPORT lean_object* lean_elab_environment_update_base_after_kernel_add(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +lean_dec(x_1); +lean_inc(x_2); +x_4 = l_Lean_Declaration_getNames(x_2); +x_5 = l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__1(x_3, x_4); +if (lean_obj_tag(x_2) == 6) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_2, 2); +lean_inc(x_6); +lean_dec(x_2); +x_7 = l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2(x_5, x_6); +return x_7; +} else { -lean_object* x_12; -lean_dec(x_3); -x_12 = l___private_Lean_Environment_0__Lean_Environment_addAux(x_1, x_2); -return x_12; +lean_dec(x_2); +return x_5; } } } @@ -18130,7 +18398,7 @@ lean_inc(x_5); x_6 = lean_ctor_get(x_4, 1); lean_inc(x_6); lean_dec(x_4); -x_7 = lean_ctor_get(x_1, 4); +x_7 = lean_ctor_get(x_1, 5); lean_inc(x_7); x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); @@ -18217,7 +18485,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean x_30 = lean_ctor_get(x_29, 1); lean_inc(x_30); lean_dec(x_29); -x_31 = lean_ctor_get(x_1, 1); +x_31 = lean_ctor_get(x_1, 0); lean_inc(x_31); x_32 = l_Lean_SMap_numBuckets___at_Lean_Environment_displayStats___spec__4(x_31); lean_dec(x_31); @@ -18246,7 +18514,7 @@ lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean x_44 = lean_ctor_get(x_43, 1); lean_inc(x_44); lean_dec(x_43); -x_45 = lean_ctor_get(x_1, 2); +x_45 = lean_ctor_get(x_1, 3); lean_inc(x_45); x_46 = lean_array_get_size(x_45); lean_dec(x_45); @@ -18662,9 +18930,8 @@ LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck___rarg(lean_object* x _start: { lean_object* x_5; -lean_inc(x_4); lean_inc(x_1); -x_5 = lean_environment_find(x_1, x_4); +x_5 = l_Lean_Environment_find_x3f(x_1, x_4); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; @@ -18750,9 +19017,7 @@ if (lean_obj_tag(x_2) == 4) { lean_object* x_3; lean_object* x_4; x_3 = lean_ctor_get(x_2, 0); -lean_inc(x_3); -lean_dec(x_2); -x_4 = lean_environment_find(x_1, x_3); +x_4 = l_Lean_Environment_find_x3f(x_1, x_3); if (lean_obj_tag(x_4) == 0) { uint8_t x_5; @@ -18773,7 +19038,6 @@ return x_7; else { uint8_t x_8; -lean_dec(x_2); lean_dec(x_1); x_8 = 0; return x_8; @@ -18808,6 +19072,7 @@ LEAN_EXPORT lean_object* l_Lean_Environment_hasUnsafe___lambda__1___boxed(lean_o { uint8_t x_3; lean_object* x_4; x_3 = l_Lean_Environment_hasUnsafe___lambda__1(x_1, x_2); +lean_dec(x_2); x_4 = lean_box(x_3); return x_4; } @@ -19156,10 +19421,18 @@ l_Lean_instInhabitedModuleData___closed__2 = _init_l_Lean_instInhabitedModuleDat lean_mark_persistent(l_Lean_instInhabitedModuleData___closed__2); l_Lean_instInhabitedModuleData = _init_l_Lean_instInhabitedModuleData(); lean_mark_persistent(l_Lean_instInhabitedModuleData); -l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__1(); -l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2 = _init_l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__2(); -l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3 = _init_l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3(); -lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3); +l_Lean_Kernel_instInhabitedDiagnostics___closed__1 = _init_l_Lean_Kernel_instInhabitedDiagnostics___closed__1(); +lean_mark_persistent(l_Lean_Kernel_instInhabitedDiagnostics___closed__1); +l_Lean_Kernel_instInhabitedDiagnostics___closed__2 = _init_l_Lean_Kernel_instInhabitedDiagnostics___closed__2(); +lean_mark_persistent(l_Lean_Kernel_instInhabitedDiagnostics___closed__2); +l_Lean_Kernel_instInhabitedDiagnostics___closed__3 = _init_l_Lean_Kernel_instInhabitedDiagnostics___closed__3(); +lean_mark_persistent(l_Lean_Kernel_instInhabitedDiagnostics___closed__3); +l_Lean_Kernel_instInhabitedDiagnostics = _init_l_Lean_Kernel_instInhabitedDiagnostics(); +lean_mark_persistent(l_Lean_Kernel_instInhabitedDiagnostics); +l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__1(); +l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Environment_find_x3f___spec__3___closed__2(); +l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__3___closed__1); l_Lean_instInhabitedEnvExtensionInterface___closed__1 = _init_l_Lean_instInhabitedEnvExtensionInterface___closed__1(); lean_mark_persistent(l_Lean_instInhabitedEnvExtensionInterface___closed__1); l_Lean_instInhabitedEnvExtensionInterface___closed__2 = _init_l_Lean_instInhabitedEnvExtensionInterface___closed__2(); @@ -19183,7 +19456,7 @@ l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1 = _init_l_Lean_E lean_mark_persistent(l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1); l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2 = _init_l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2(); lean_mark_persistent(l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2); -if (builtin) {res = l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1472_(lean_io_mk_world()); +if (builtin) {res = l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1833_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef = lean_io_result_get_value(res); lean_mark_persistent(l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_envExtensionsRef); @@ -19238,10 +19511,6 @@ l_Lean_mkEmptyEnvironment___lambda__1___closed__3 = _init_l_Lean_mkEmptyEnvironm lean_mark_persistent(l_Lean_mkEmptyEnvironment___lambda__1___closed__3); l_Lean_mkEmptyEnvironment___lambda__1___closed__4 = _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__4(); lean_mark_persistent(l_Lean_mkEmptyEnvironment___lambda__1___closed__4); -l_Lean_mkEmptyEnvironment___lambda__1___closed__5 = _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__5(); -lean_mark_persistent(l_Lean_mkEmptyEnvironment___lambda__1___closed__5); -l_Lean_mkEmptyEnvironment___lambda__1___closed__6 = _init_l_Lean_mkEmptyEnvironment___lambda__1___closed__6(); -lean_mark_persistent(l_Lean_mkEmptyEnvironment___lambda__1___closed__6); l_Lean_mkEmptyEnvironment___closed__1 = _init_l_Lean_mkEmptyEnvironment___closed__1(); lean_mark_persistent(l_Lean_mkEmptyEnvironment___closed__1); l_Lean_mkEmptyEnvironment___closed__2 = _init_l_Lean_mkEmptyEnvironment___closed__2(); @@ -19256,69 +19525,69 @@ l_Lean_instInhabitedPersistentEnvExtension___closed__4 = _init_l_Lean_instInhabi lean_mark_persistent(l_Lean_instInhabitedPersistentEnvExtension___closed__4); l_Lean_instInhabitedPersistentEnvExtension___closed__5 = _init_l_Lean_instInhabitedPersistentEnvExtension___closed__5(); lean_mark_persistent(l_Lean_instInhabitedPersistentEnvExtension___closed__5); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_2639_(lean_io_mk_world()); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_3055_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_persistentEnvExtensionsRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_persistentEnvExtensionsRef); lean_dec_ref(res); -}l___auto____x40_Lean_Environment___hyg_2684____closed__1 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__1(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__1); -l___auto____x40_Lean_Environment___hyg_2684____closed__2 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__2(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__2); -l___auto____x40_Lean_Environment___hyg_2684____closed__3 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__3(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__3); -l___auto____x40_Lean_Environment___hyg_2684____closed__4 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__4(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__4); -l___auto____x40_Lean_Environment___hyg_2684____closed__5 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__5(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__5); -l___auto____x40_Lean_Environment___hyg_2684____closed__6 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__6(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__6); -l___auto____x40_Lean_Environment___hyg_2684____closed__7 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__7(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__7); -l___auto____x40_Lean_Environment___hyg_2684____closed__8 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__8(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__8); -l___auto____x40_Lean_Environment___hyg_2684____closed__9 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__9(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__9); -l___auto____x40_Lean_Environment___hyg_2684____closed__10 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__10(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__10); -l___auto____x40_Lean_Environment___hyg_2684____closed__11 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__11(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__11); -l___auto____x40_Lean_Environment___hyg_2684____closed__12 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__12(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__12); -l___auto____x40_Lean_Environment___hyg_2684____closed__13 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__13(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__13); -l___auto____x40_Lean_Environment___hyg_2684____closed__14 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__14(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__14); -l___auto____x40_Lean_Environment___hyg_2684____closed__15 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__15(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__15); -l___auto____x40_Lean_Environment___hyg_2684____closed__16 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__16(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__16); -l___auto____x40_Lean_Environment___hyg_2684____closed__17 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__17(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__17); -l___auto____x40_Lean_Environment___hyg_2684____closed__18 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__18(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__18); -l___auto____x40_Lean_Environment___hyg_2684____closed__19 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__19(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__19); -l___auto____x40_Lean_Environment___hyg_2684____closed__20 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__20(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__20); -l___auto____x40_Lean_Environment___hyg_2684____closed__21 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__21(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__21); -l___auto____x40_Lean_Environment___hyg_2684____closed__22 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__22(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__22); -l___auto____x40_Lean_Environment___hyg_2684____closed__23 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__23(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__23); -l___auto____x40_Lean_Environment___hyg_2684____closed__24 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__24(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__24); -l___auto____x40_Lean_Environment___hyg_2684____closed__25 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__25(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__25); -l___auto____x40_Lean_Environment___hyg_2684____closed__26 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__26(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__26); -l___auto____x40_Lean_Environment___hyg_2684____closed__27 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__27(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__27); -l___auto____x40_Lean_Environment___hyg_2684____closed__28 = _init_l___auto____x40_Lean_Environment___hyg_2684____closed__28(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684____closed__28); -l___auto____x40_Lean_Environment___hyg_2684_ = _init_l___auto____x40_Lean_Environment___hyg_2684_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2684_); +}l___auto____x40_Lean_Environment___hyg_3100____closed__1 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__1(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__1); +l___auto____x40_Lean_Environment___hyg_3100____closed__2 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__2(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__2); +l___auto____x40_Lean_Environment___hyg_3100____closed__3 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__3(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__3); +l___auto____x40_Lean_Environment___hyg_3100____closed__4 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__4(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__4); +l___auto____x40_Lean_Environment___hyg_3100____closed__5 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__5(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__5); +l___auto____x40_Lean_Environment___hyg_3100____closed__6 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__6(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__6); +l___auto____x40_Lean_Environment___hyg_3100____closed__7 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__7(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__7); +l___auto____x40_Lean_Environment___hyg_3100____closed__8 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__8(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__8); +l___auto____x40_Lean_Environment___hyg_3100____closed__9 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__9(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__9); +l___auto____x40_Lean_Environment___hyg_3100____closed__10 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__10(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__10); +l___auto____x40_Lean_Environment___hyg_3100____closed__11 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__11(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__11); +l___auto____x40_Lean_Environment___hyg_3100____closed__12 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__12(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__12); +l___auto____x40_Lean_Environment___hyg_3100____closed__13 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__13(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__13); +l___auto____x40_Lean_Environment___hyg_3100____closed__14 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__14(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__14); +l___auto____x40_Lean_Environment___hyg_3100____closed__15 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__15(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__15); +l___auto____x40_Lean_Environment___hyg_3100____closed__16 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__16(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__16); +l___auto____x40_Lean_Environment___hyg_3100____closed__17 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__17(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__17); +l___auto____x40_Lean_Environment___hyg_3100____closed__18 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__18(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__18); +l___auto____x40_Lean_Environment___hyg_3100____closed__19 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__19(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__19); +l___auto____x40_Lean_Environment___hyg_3100____closed__20 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__20(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__20); +l___auto____x40_Lean_Environment___hyg_3100____closed__21 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__21(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__21); +l___auto____x40_Lean_Environment___hyg_3100____closed__22 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__22(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__22); +l___auto____x40_Lean_Environment___hyg_3100____closed__23 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__23(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__23); +l___auto____x40_Lean_Environment___hyg_3100____closed__24 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__24(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__24); +l___auto____x40_Lean_Environment___hyg_3100____closed__25 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__25(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__25); +l___auto____x40_Lean_Environment___hyg_3100____closed__26 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__26(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__26); +l___auto____x40_Lean_Environment___hyg_3100____closed__27 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__27(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__27); +l___auto____x40_Lean_Environment___hyg_3100____closed__28 = _init_l___auto____x40_Lean_Environment___hyg_3100____closed__28(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100____closed__28); +l___auto____x40_Lean_Environment___hyg_3100_ = _init_l___auto____x40_Lean_Environment___hyg_3100_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3100_); l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1 = _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1(); lean_mark_persistent(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1); l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2 = _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2(); @@ -19327,16 +19596,16 @@ l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1 = _init_l_Lean_re lean_mark_persistent(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1); l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2 = _init_l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2(); lean_mark_persistent(l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__2); -l___auto____x40_Lean_Environment___hyg_3062_ = _init_l___auto____x40_Lean_Environment___hyg_3062_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3062_); +l___auto____x40_Lean_Environment___hyg_3478_ = _init_l___auto____x40_Lean_Environment___hyg_3478_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3478_); l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1 = _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1(); lean_mark_persistent(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__1); l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2 = _init_l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2(); lean_mark_persistent(l_Lean_registerSimplePersistentEnvExtension___rarg___lambda__4___closed__2); l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1 = _init_l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1(); lean_mark_persistent(l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1); -l___auto____x40_Lean_Environment___hyg_3379_ = _init_l___auto____x40_Lean_Environment___hyg_3379_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3379_); +l___auto____x40_Lean_Environment___hyg_3795_ = _init_l___auto____x40_Lean_Environment___hyg_3795_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3795_); l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1 = _init_l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1(); lean_mark_persistent(l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1___closed__1); l_Lean_mkTagDeclarationExtension___closed__1 = _init_l_Lean_mkTagDeclarationExtension___closed__1(); @@ -19361,8 +19630,8 @@ l_Lean_TagDeclarationExtension_tag___closed__5 = _init_l_Lean_TagDeclarationExte lean_mark_persistent(l_Lean_TagDeclarationExtension_tag___closed__5); l_Lean_TagDeclarationExtension_isTagged___closed__1 = _init_l_Lean_TagDeclarationExtension_isTagged___closed__1(); lean_mark_persistent(l_Lean_TagDeclarationExtension_isTagged___closed__1); -l___auto____x40_Lean_Environment___hyg_3542_ = _init_l___auto____x40_Lean_Environment___hyg_3542_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3542_); +l___auto____x40_Lean_Environment___hyg_3958_ = _init_l___auto____x40_Lean_Environment___hyg_3958_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3958_); l_Lean_mkMapDeclarationExtension___rarg___closed__1 = _init_l_Lean_mkMapDeclarationExtension___rarg___closed__1(); lean_mark_persistent(l_Lean_mkMapDeclarationExtension___rarg___closed__1); l_Lean_mkMapDeclarationExtension___rarg___closed__2 = _init_l_Lean_mkMapDeclarationExtension___rarg___closed__2(); @@ -19419,46 +19688,35 @@ 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_6600____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__2); -l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____lambda__2___closed__3); -l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__2); -l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__3); -l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__4); -l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__5); -l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6600____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6600_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__2); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____lambda__2___closed__3); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__2); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__3); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__4); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__5); +l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_7023____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_7023_(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); lean_dec_ref(res); -}l_Lean_Kernel_instInhabitedDiagnostics___closed__1 = _init_l_Lean_Kernel_instInhabitedDiagnostics___closed__1(); -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_6727____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6727____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6727____closed__1); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6727_(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); -lean_dec_ref(res); -}l_Lean_Kernel_enableDiag___closed__1 = _init_l_Lean_Kernel_enableDiag___closed__1(); -lean_mark_persistent(l_Lean_Kernel_enableDiag___closed__1); -l_Lean_Kernel_resetDiag___closed__1 = _init_l_Lean_Kernel_resetDiag___closed__1(); -lean_mark_persistent(l_Lean_Kernel_resetDiag___closed__1); -l_Lean_Environment_registerNamespace___closed__1 = _init_l_Lean_Environment_registerNamespace___closed__1(); +}l_Lean_Environment_registerNamespace___closed__1 = _init_l_Lean_Environment_registerNamespace___closed__1(); lean_mark_persistent(l_Lean_Environment_registerNamespace___closed__1); +l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__1 = _init_l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__1(); +lean_mark_persistent(l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__1); +l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__2 = _init_l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__2(); +lean_mark_persistent(l_List_foldl___at___private_Lean_Environment_0__Lean_Environment_updateBaseAfterKernelAdd___spec__2___closed__2); l_List_foldl___at_Lean_Environment_displayStats___spec__2___closed__1 = _init_l_List_foldl___at_Lean_Environment_displayStats___spec__2___closed__1(); lean_mark_persistent(l_List_foldl___at_Lean_Environment_displayStats___spec__2___closed__1); l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__1 = _init_l_List_toString___at_Lean_Environment_displayStats___spec__1___closed__1(); diff --git a/stage0/stdlib/Lean/Exception.c b/stage0/stdlib/Lean/Exception.c index 03503c51efe8..9b6020f7d0ee 100644 --- a/stage0/stdlib/Lean/Exception.c +++ b/stage0/stdlib/Lean/Exception.c @@ -119,6 +119,7 @@ static lean_object* l_Lean_termThrowError_______closed__1; LEAN_EXPORT lean_object* l_Lean_Exception_isMaxRecDepth___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withIncRecDepth___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_termThrowErrorAt___________closed__8; @@ -152,7 +153,6 @@ static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThr lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__1; LEAN_EXPORT lean_object* l_Lean_withIncRecDepth(lean_object*, lean_object*); -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); lean_object* l_String_toSubstring_x27(lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; static lean_object* l_Lean_termThrowError_______closed__9; @@ -536,7 +536,7 @@ LEAN_EXPORT lean_object* l_Lean_throwKernelException___rarg___lambda__1(lean_obj _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_KernelException_toMessageData(x_1, x_4); +x_5 = l_Lean_Kernel_Exception_toMessageData(x_1, x_4); x_6 = l_Lean_throwError___rarg(x_2, x_3, x_5); return x_6; } diff --git a/stage0/stdlib/Lean/Language/Lean.c b/stage0/stdlib/Lean/Language/Lean.c index 14f0fdf2f54d..138d863f09e3 100644 --- a/stage0/stdlib/Lean/Language/Lean.c +++ b/stage0/stdlib/Lean/Language/Lean.c @@ -119,8 +119,6 @@ LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader(lean_object*, LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__7; lean_object* l_Lean_Elab_withLogging___at_Lean_Elab_Command_withLoggingExceptions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__17; static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5; static lean_object* l_List_forIn_x27_loop___at_Lean_Language_Lean_reparseOptions___spec__1___lambda__1___closed__3; @@ -161,6 +159,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Language_Lean_pr static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__12___closed__13; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_402____closed__4; lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instMonadLiftProcessingTLeanProcessingT___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__1(lean_object*); @@ -195,7 +194,7 @@ lean_object* lean_io_mono_nanos_now(lean_object*); lean_object* l_panic___at_String_fromUTF8_x21___spec__1(lean_object*); static lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__6___closed__2; static lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg___closed__2; -lean_object* lean_environment_set_main_module(lean_object*, lean_object*); +lean_object* l_Lean_Environment_setMainModule(lean_object*, lean_object*); lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); lean_object* lean_thunk_get_own(lean_object*); static lean_object* l_Lean_Language_Lean_process_doElab___closed__4; @@ -342,6 +341,7 @@ static lean_object* l_List_forIn_x27_loop___at_Lean_Language_Lean_reparseOptions LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Language_Lean_process_parseCmd___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_402_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Language_Lean_process_parseCmd___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_402____closed__6; lean_object* l_Lean_Parser_parseHeader(lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_process_doElab___lambda__4___closed__3; @@ -5881,7 +5881,7 @@ lean_dec(x_83); x_86 = lean_ctor_get(x_84, 0); lean_inc(x_86); lean_dec(x_84); -x_87 = l_Lean_Kernel_isDiagnosticsEnabled(x_86); +x_87 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_86); lean_dec(x_86); if (x_87 == 0) { @@ -6118,7 +6118,7 @@ lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean x_93 = lean_ctor_get(x_90, 0); x_94 = lean_ctor_get(x_90, 4); lean_dec(x_94); -x_95 = l_Lean_Kernel_enableDiag(x_93, x_82); +x_95 = l_Lean_Kernel_Environment_enableDiag(x_93, x_82); lean_ctor_set(x_90, 4, x_76); lean_ctor_set(x_90, 0, x_95); x_96 = lean_st_ref_set(x_79, x_90, x_91); @@ -6204,7 +6204,7 @@ lean_inc(x_116); lean_inc(x_115); lean_inc(x_114); lean_dec(x_90); -x_121 = l_Lean_Kernel_enableDiag(x_114, x_82); +x_121 = l_Lean_Kernel_Environment_enableDiag(x_114, x_82); x_122 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_122, 0, x_121); lean_ctor_set(x_122, 1, x_115); @@ -7092,7 +7092,7 @@ lean_dec(x_83); x_86 = lean_ctor_get(x_84, 0); lean_inc(x_86); lean_dec(x_84); -x_87 = l_Lean_Kernel_isDiagnosticsEnabled(x_86); +x_87 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_86); lean_dec(x_86); if (x_87 == 0) { @@ -7329,7 +7329,7 @@ lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean x_93 = lean_ctor_get(x_90, 0); x_94 = lean_ctor_get(x_90, 4); lean_dec(x_94); -x_95 = l_Lean_Kernel_enableDiag(x_93, x_82); +x_95 = l_Lean_Kernel_Environment_enableDiag(x_93, x_82); lean_ctor_set(x_90, 4, x_76); lean_ctor_set(x_90, 0, x_95); x_96 = lean_st_ref_set(x_79, x_90, x_91); @@ -7415,7 +7415,7 @@ lean_inc(x_116); lean_inc(x_115); lean_inc(x_114); lean_dec(x_90); -x_121 = l_Lean_Kernel_enableDiag(x_114, x_82); +x_121 = l_Lean_Kernel_Environment_enableDiag(x_114, x_82); x_122 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_122, 0, x_121); lean_ctor_set(x_122, 1, x_115); @@ -12945,7 +12945,7 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_15 = lean_ctor_get(x_1, 0); lean_inc(x_15); lean_dec(x_1); -x_16 = lean_environment_set_main_module(x_2, x_15); +x_16 = l_Lean_Environment_setMainModule(x_2, x_15); x_17 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2; x_18 = l_Lean_Option_get_x3f___at_Lean_addTraceAsMessages___spec__17(x_3, x_17); if (lean_obj_tag(x_18) == 0) diff --git a/stage0/stdlib/Lean/Linter/ConstructorAsVariable.c b/stage0/stdlib/Lean/Linter/ConstructorAsVariable.c index b56b6c6d885d..9d4ced534ad2 100644 --- a/stage0/stdlib/Lean/Linter/ConstructorAsVariable.c +++ b/stage0/stdlib/Lean/Linter/ConstructorAsVariable.c @@ -47,7 +47,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_constructorNam LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_constructorNameAsVariable___elambda__1___spec__14(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_linter_constructorNameAsVariable; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Linter_constructorNameAsVariable___elambda__1___spec__5(lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_constructorNameAsVariable___elambda__1___spec__14___closed__9; lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_constructorNameAsVariable___elambda__1___spec__8___closed__1; @@ -1024,8 +1024,7 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -lean_inc(x_17); -x_30 = lean_environment_find(x_29, x_17); +x_30 = l_Lean_Environment_find_x3f(x_29, x_17); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_object* x_32; @@ -1107,8 +1106,7 @@ lean_dec(x_25); x_45 = lean_ctor_get(x_43, 0); lean_inc(x_45); lean_dec(x_43); -lean_inc(x_17); -x_46 = lean_environment_find(x_45, x_17); +x_46 = l_Lean_Environment_find_x3f(x_45, x_17); if (lean_obj_tag(x_46) == 0) { lean_object* x_47; lean_object* x_48; @@ -2058,7 +2056,8 @@ x_23 = lean_ctor_get(x_20, 1); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_find(x_24, x_19); +x_25 = l_Lean_Environment_find_x3f(x_24, x_19); +lean_dec(x_19); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; @@ -2137,7 +2136,8 @@ lean_dec(x_20); x_40 = lean_ctor_get(x_38, 0); lean_inc(x_40); lean_dec(x_38); -x_41 = lean_environment_find(x_40, x_19); +x_41 = l_Lean_Environment_find_x3f(x_40, x_19); +lean_dec(x_19); if (lean_obj_tag(x_41) == 0) { lean_object* x_42; lean_object* x_43; @@ -2252,7 +2252,8 @@ if (lean_is_exclusive(x_60)) { x_64 = lean_ctor_get(x_61, 0); lean_inc(x_64); lean_dec(x_61); -x_65 = lean_environment_find(x_64, x_59); +x_65 = l_Lean_Environment_find_x3f(x_64, x_59); +lean_dec(x_59); if (lean_obj_tag(x_65) == 0) { lean_object* x_66; lean_object* x_67; diff --git a/stage0/stdlib/Lean/Linter/MissingDocs.c b/stage0/stdlib/Lean/Linter/MissingDocs.c index 83759bb8497e..e82632ecfbbf 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Linter.MissingDocs -// Imports: Lean.Meta.Tactic.Simp.RegisterCommand Lean.Elab.Command Lean.Elab.SetOption Lean.Linter.Util +// Imports: Lean.Parser.Syntax Lean.Meta.Tactic.Simp.RegisterCommand Lean.Elab.Command Lean.Elab.SetOption Lean.Linter.Util #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -102,7 +102,7 @@ static lean_object* l_Lean_Linter_MissingDocs_missingDocs___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_mkSimpleHandler___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkSyntax(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___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_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__2; static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_802____lambda__2___closed__1; static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_802____closed__14; @@ -815,9 +815,8 @@ lean_inc(x_4); x_5 = lean_ctor_get(x_2, 1); lean_inc(x_5); lean_dec(x_2); -lean_inc(x_1); lean_inc(x_4); -x_6 = lean_environment_find(x_4, x_1); +x_6 = l_Lean_Environment_find_x3f(x_4, x_1); if (lean_obj_tag(x_6) == 0) { uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; @@ -8777,6 +8776,7 @@ x_4 = l_Lean_Linter_MissingDocs_addBuiltinHandler(x_2, x_3, x_1); return x_4; } } +lean_object* initialize_Lean_Parser_Syntax(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Simp_RegisterCommand(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Command(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_SetOption(uint8_t builtin, lean_object*); @@ -8786,6 +8786,9 @@ LEAN_EXPORT lean_object* initialize_Lean_Linter_MissingDocs(uint8_t builtin, lea lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; +res = initialize_Lean_Parser_Syntax(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Simp_RegisterCommand(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Linter/UnusedVariables.c b/stage0/stdlib/Lean/Linter/UnusedVariables.c index 8862d98debe7..da45d506c444 100644 --- a/stage0/stdlib/Lean/Linter/UnusedVariables.c +++ b/stage0/stdlib/Lean/Linter/UnusedVariables.c @@ -141,7 +141,7 @@ LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_UnusedVariables_ LEAN_EXPORT lean_object* l_Lean_Linter_getLinterUnusedVariablesFunArgs___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_UnusedVariables_unusedVariables___elambda__1___spec__6___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_987____lambda__1___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_640____lambda__1___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1316____lambda__4___closed__10; lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1242,9 +1242,8 @@ x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); x_5 = lean_ctor_get(x_2, 1); lean_inc(x_5); -lean_inc(x_1); lean_inc(x_4); -x_6 = lean_environment_find(x_4, x_1); +x_6 = l_Lean_Environment_find_x3f(x_4, x_1); if (lean_obj_tag(x_6) == 0) { uint8_t x_7; diff --git a/stage0/stdlib/Lean/Message.c b/stage0/stdlib/Lean/Message.c index a36217534d69..1feea10414cf 100644 --- a/stage0/stdlib/Lean/Message.c +++ b/stage0/stdlib/Lean/Message.c @@ -19,6 +19,7 @@ LEAN_EXPORT uint8_t l_Lean_instInhabitedMessageSeverity; LEAN_EXPORT lean_object* l_Lean_instToMessageDataOptionExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_toArray___boxed(lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__16; static lean_object* l_Lean_MessageData_instCoeExpr___closed__1; LEAN_EXPORT lean_object* l_Lean_Message_toString___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MessageData_lazy___elambda__1___closed__2; @@ -31,8 +32,10 @@ LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString___lambda__2(lean_object*, static lean_object* l_Lean_SerialMessage_toString___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_paren(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2737____rarg___closed__8; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__53; extern lean_object* l_Std_instInhabitedFormat; static lean_object* l_Lean_mkErrorStringWithPos___closed__4; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__30; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__32; LEAN_EXPORT lean_object* l_Lean_MessageData_kind(lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); @@ -41,16 +44,17 @@ static lean_object* l_Lean_MessageData_formatAux___lambda__1___closed__1; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__3; lean_object* lean_mk_empty_array_with_capacity(lean_object*); size_t lean_usize_shift_right(size_t, size_t); -static lean_object* l_Lean_KernelException_toMessageData___closed__4; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__8; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____spec__1___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__5; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__52; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__28; static lean_object* l_List_mapTR_loop___at_Lean_MessageData_orList___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_SerialMessage_toMessage(lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__50; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__3; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__16; -static lean_object* l_Lean_KernelException_toMessageData___closed__36; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__39; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__41; lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_MessageLog_errorsToWarnings___spec__2(lean_object*); lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); @@ -58,7 +62,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageData_formatAux_ static lean_object* l_Lean_SerialMessage_toString___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_MessageLog_toList___boxed(lean_object*); static lean_object* l_Lean_MessageData_orList___closed__5; -static lean_object* l_Lean_KernelException_toMessageData___closed__6; LEAN_EXPORT lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__24; @@ -66,7 +69,7 @@ LEAN_EXPORT lean_object* l_Lean_indentD(lean_object*); static lean_object* l_Lean_termM_x21_____closed__1; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_MessageData_lazy___elambda__1___closed__1; -static lean_object* l_Lean_KernelException_toMessageData___closed__19; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__32; LEAN_EXPORT lean_object* l_Lean_MessageData_ofSyntax___elambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__2; @@ -75,6 +78,7 @@ lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static size_t l_Lean_instInhabitedMessageLog___closed__3; static lean_object* l_Lean_MessageData_ofList___closed__4; static lean_object* l_Lean_MessageData_orList___closed__3; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__33; LEAN_EXPORT lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_Lean_MessageData_arrayExpr_toMessageData___closed__3; lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); @@ -88,13 +92,12 @@ static lean_object* l_Lean_instToMessageDataOption___rarg___closed__2; static lean_object* l_List_mapTR_loop___at_Lean_MessageData_orList___spec__1___closed__1; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__31; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__8; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__26; uint8_t lean_usize_dec_eq(size_t, size_t); -static lean_object* l_Lean_KernelException_toMessageData___closed__37; static lean_object* l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2737____rarg___closed__7; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__12; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__31; static lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____lambda__2___closed__1; uint8_t l_Lean_Name_isAnonymous(lean_object*); static lean_object* l_Lean_MessageData_andList___closed__2; @@ -105,20 +108,21 @@ static lean_object* l_Lean_MessageData_hasSyntheticSorry_visit___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_MessageLog_errorsToWarnings___spec__1(lean_object*); static lean_object* l_Lean_MessageData_hasSyntheticSorry_visit___closed__1; lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__16; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeFormat; static lean_object* l_Lean_MessageData_formatAux___lambda__1___closed__2; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__36; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToMessageDataOption___rarg(lean_object*, lean_object*); static double l_Lean_MessageData_formatAux___closed__5; LEAN_EXPORT lean_object* l_String_splitAux___at_Lean_stringToMessageData___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__45; LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____lambda__1(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__3___closed__1; static lean_object* l_Lean_MessageData_ofList___closed__7; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__34; LEAN_EXPORT lean_object* l_Lean_instFromJsonSerialMessage; LEAN_EXPORT lean_object* l_Lean_MessageData_hasTag___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeArrayExpr(lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__25; static lean_object* l_Lean_MessageData_instCoeOptionExpr___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MessageData_hasTag___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__15; @@ -130,7 +134,9 @@ LEAN_EXPORT lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_107____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_formatAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedBaseMessage___rarg___closed__1; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__19; LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____lambda__1___boxed(lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__8; LEAN_EXPORT lean_object* l_Lean_MessageData_formatAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__22; LEAN_EXPORT lean_object* l_Lean_mkErrorStringWithPos(lean_object*, lean_object*, lean_object*, lean_object*); @@ -140,7 +146,6 @@ LEAN_EXPORT lean_object* l_Lean_instToMessageDataSyntax; LEAN_EXPORT lean_object* l_Lean_instAddMessageContextOfMonadLift___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____spec__1___closed__3; static lean_object* l_Lean_termM_x21_____closed__12; -static lean_object* l_Lean_KernelException_toMessageData___closed__34; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_instBEqMessageSeverity___closed__1; @@ -149,19 +154,20 @@ static lean_object* l_Lean_MessageLog_append___closed__1; static lean_object* l_Lean_instImpl____x40_Lean_Message___hyg_606____closed__2; lean_object* l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__1; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__13; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeListExpr(lean_object*); static lean_object* l_Lean_MessageData_orList___closed__2; lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); static lean_object* l_Lean_MessageData_formatAux___lambda__1___closed__6; lean_object* l_Lean_PersistentArray_toList___rarg(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2737____rarg___closed__6; -static lean_object* l_Lean_KernelException_toMessageData___closed__32; static lean_object* l_Lean_MessageData_formatAux___closed__3; static lean_object* l_Lean_MessageLog_empty___closed__1; lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Message_toJson(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_lazy___elambda__1___closed__3; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__40; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeString; static lean_object* l_Lean_MessageData_nil___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_instAppend(lean_object*, lean_object*); @@ -169,6 +175,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax___rarg(lean_object*); static lean_object* l_Lean_MessageData_orList___closed__4; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____spec__1___closed__2; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__42; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeSyntax; LEAN_EXPORT lean_object* l_Lean_instToMessageDataOfToFormat___rarg(lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); @@ -180,24 +187,31 @@ static lean_object* l_Lean_MessageData_instCoeString___closed__3; LEAN_EXPORT lean_object* l_Lean_addMessageContextFull(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedBaseMessage___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_ofExpr___elambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__43; uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeOptionExpr(lean_object*); static lean_object* l_Lean_instInhabitedMessageData___closed__1; lean_object* l_List_mapTR_loop___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__27; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__19; -static lean_object* l_Lean_KernelException_toMessageData___closed__44; static lean_object* l_Lean_MessageData_instCoeList___closed__1; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__50; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__6; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__4; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__29; lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__6; static lean_object* l_Lean_mkErrorStringWithPos___closed__3; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeString___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_instToMessageDataArray___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__24; static lean_object* l_Lean_MessageData_orList___closed__1; LEAN_EXPORT uint8_t l_Lean_MessageData_ofConstName___elambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_reportedPlusUnreported(lean_object*); static lean_object* l_Lean_instFromJsonMessageSeverity___closed__1; static lean_object* l_Lean_termM_x21_____closed__7; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__35; lean_object* l_Lean_Json_getObjValAs_x3f(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_MessageData_instCoeArrayExpr___closed__2; @@ -207,8 +221,6 @@ static lean_object* l_Lean_MessageData_ofList___closed__6; static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2737____rarg___closed__2; lean_object* l_List_getLast_x21___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____closed__1; -static lean_object* l_Lean_KernelException_toMessageData___closed__1; -static lean_object* l_Lean_KernelException_toMessageData___closed__30; lean_object* l_Lean_Syntax_copyHeadTailInfoFrom(lean_object*, lean_object*); lean_object* l_Lean_formatRawGoal(lean_object*); static lean_object* l_Lean_MessageData_formatAux___closed__8; @@ -216,10 +228,9 @@ lean_object* l_Lean_ppLevel(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MessageData_formatAux___closed__4; LEAN_EXPORT lean_object* l_String_split___at_Lean_stringToMessageData___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1___boxed(lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__35; -static lean_object* l_Lean_KernelException_toMessageData___closed__38; lean_object* lean_string_utf8_next(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeLevel; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__14; static lean_object* l_Lean_MessageData_arrayExpr_toMessageData___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____spec__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_instToMessageDataOption___rarg___closed__3; @@ -242,14 +253,12 @@ LEAN_EXPORT lean_object* l_Lean_instToMessageDataFormat; LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_MessageLog_getInfoMessages___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedMessageData; lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__3; static lean_object* l_Lean_MessageLog_instAppend___closed__1; LEAN_EXPORT lean_object* l_Lean_instToMessageDataSubarray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_toList(lean_object*); lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_ofList___closed__5; -static lean_object* l_Lean_KernelException_toMessageData___closed__33; -static lean_object* l_Lean_KernelException_toMessageData___closed__10; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__26; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2737____rarg(lean_object*, lean_object*); @@ -257,7 +266,6 @@ static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_ static lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____lambda__2___closed__2; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__13; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MessageLog_getInfoMessages___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__24; LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_formatAux___closed__6; @@ -266,8 +274,6 @@ static lean_object* l_Lean_MessageData_instCoeLevel___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_andList(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_toJsonMessageSeverity____x40_Lean_Message___hyg_125____closed__2; static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__11; -static lean_object* l_Lean_KernelException_toMessageData___closed__45; -static lean_object* l_Lean_KernelException_toMessageData___closed__46; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__27; LEAN_EXPORT lean_object* l_Lean_MessageData_ofExpr___elambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqMessageSeverity; @@ -275,9 +281,9 @@ LEAN_EXPORT lean_object* l_Lean_Message_kind(lean_object*); static lean_object* l_Lean_MessageData_ofConstName___elambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_instTypeNameMessageData; lean_object* lean_array_pop(lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__38; LEAN_EXPORT lean_object* l_Lean_MessageSeverity_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedMessageLog; -static lean_object* l_Lean_KernelException_toMessageData___closed__9; LEAN_EXPORT lean_object* l_Lean_instToMessageDataString; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__5; lean_object* lean_array_to_list(lean_object*); @@ -286,16 +292,15 @@ static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArra static lean_object* l_Lean_instImpl____x40_Lean_Message___hyg_606____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357_(lean_object*); LEAN_EXPORT uint8_t l_Lean_MessageData_hasTag(lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__22; static lean_object* l_Lean_instToMessageDataOption___rarg___closed__4; static lean_object* l___private_Lean_Message_0__Lean_toJsonMessageSeverity____x40_Lean_Message___hyg_125____closed__5; -static lean_object* l_Lean_KernelException_toMessageData___closed__29; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__4(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_ofList___closed__2; static lean_object* l_Lean_MessageData_instCoeSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageSeverity_noConfusion(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__25; -static lean_object* l_Lean_KernelException_toMessageData___closed__2; uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); static lean_object* l_Lean_instToMessageDataMessageData___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_nil; @@ -310,6 +315,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_MessageLog_hasError static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__14; LEAN_EXPORT lean_object* l_Lean_instFromJsonBaseMessage(lean_object*); lean_object* l_Lean_ppConstNameWithInfos(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__12; static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2737____rarg___closed__9; LEAN_EXPORT lean_object* l_Lean_MessageData_orList(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__4; @@ -319,7 +325,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageData_ofConstName___boxed(lean_object*, le lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__15; LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__12; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__17; static lean_object* l_Lean_termM_x21_____closed__4; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__11; static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__10; @@ -340,7 +346,6 @@ lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_ static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2737____rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_ofLevel___elambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_split___at_Lean_stringToMessageData___spec__1___boxed(lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__53; LEAN_EXPORT lean_object* l_Lean_instToMessageDataProd___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_MessageLog_getInfoMessages___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString(lean_object*, uint8_t); @@ -351,15 +356,19 @@ LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString___lambda__1(lean_object*, extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_MessageLog_getInfoMessages___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_arrayExpr_toMessageData(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__51; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__8; LEAN_EXPORT lean_object* l_Lean_instFromJsonMessageSeverity; LEAN_EXPORT lean_object* l_Lean_MessageData_ofName___lambda__1___boxed(lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__25; lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__29; -static lean_object* l_Lean_KernelException_toMessageData___closed__40; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__44; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__23; LEAN_EXPORT lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__13; LEAN_EXPORT lean_object* l_Lean_MessageSeverity_toCtorIdx(uint8_t); @@ -379,24 +388,20 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors__ double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__7; static lean_object* l_Lean_instToMessageDataOptionExpr___closed__3; -static lean_object* l_Lean_KernelException_toMessageData___closed__7; static lean_object* l_Lean_instToMessageDataOptionExpr___closed__1; static lean_object* l_panic___at_Lean_MessageData_formatAux___spec__2___closed__1; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__23; static lean_object* l_Lean_termM_x21_____closed__8; -static lean_object* l_Lean_KernelException_toMessageData___closed__41; LEAN_EXPORT lean_object* l_Lean_instToMessageDataProd(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164_(lean_object*); LEAN_EXPORT lean_object* l_Lean_instToMessageDataMVarId(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__3; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__17; -static lean_object* l_Lean_KernelException_toMessageData___closed__43; lean_object* l___private_Lean_Data_Position_0__Lean_fromJsonPosition____x40_Lean_Data_Position___hyg_289_(lean_object*); LEAN_EXPORT uint8_t l_Lean_MessageData_ofName___lambda__1(lean_object*); static lean_object* l_Lean_toMessageList___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_mkPPContext(lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Lean_Message___hyg_2737____rarg___closed__5; -static lean_object* l_Lean_KernelException_toMessageData___closed__42; static lean_object* l_Lean_MessageData_formatAux___closed__9; LEAN_EXPORT lean_object* l_Lean_MessageSeverity_noConfusion___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instToMessageDataList(lean_object*); @@ -405,13 +410,13 @@ lean_object* l_Lean_ppTerm(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instAddMessageContextOfMonadLift(lean_object*, lean_object*); lean_object* l_Lean_Json_parseTagged(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__5; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__1; static lean_object* l_Lean_instToMessageDataOption___rarg___closed__5; static lean_object* l_Lean_MessageData_arrayExpr_toMessageData___closed__4; static lean_object* l_Lean_MessageData_formatAux___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_MessageData_lazy___elambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__9; static lean_object* l_Lean_MessageSeverity_noConfusion___rarg___closed__1; -static lean_object* l_Lean_KernelException_toMessageData___closed__15; extern lean_object* l_Lean_NameSet_empty; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__17; static lean_object* l___private_Lean_Message_0__Lean_toJsonMessageSeverity____x40_Lean_Message___hyg_125____closed__4; @@ -424,15 +429,16 @@ static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Le lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MessageData_hasSyntheticSorry_visit___spec__1(lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__28; static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__3; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__2; static lean_object* l_Lean_MessageData_arrayExpr_toMessageData___closed__2; -static lean_object* l_Lean_KernelException_toMessageData___closed__48; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____lambda__3(lean_object*, lean_object*, lean_object*); lean_object* lean_float_to_string(double); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverity____x40_Lean_Message___hyg_164____lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_ofLevel(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__9; @@ -441,11 +447,14 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message LEAN_EXPORT lean_object* l_Lean_MessageLog_getInfoMessages(lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_forM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Message_serialize(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__4; LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString___lambda__2___boxed(lean_object*, lean_object*); lean_object* l_id___rarg___boxed(lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__46; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__11; static lean_object* l_Lean_MessageData_ofLevel___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_toJsonMessageSeverity____x40_Lean_Message___hyg_125____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_termM_x21__; @@ -457,32 +466,29 @@ static lean_object* l_Lean_MessageData_andList___closed__3; LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MessageLog_empty___closed__3; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__5(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__28; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__15; LEAN_EXPORT uint8_t l_Lean_MessageData_ofExpr___elambda__1(lean_object*, lean_object*); lean_object* l_Std_Format_joinSep___at_Prod_repr___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_SerialMessage_toString___closed__1; -static lean_object* l_Lean_KernelException_toMessageData___closed__39; uint8_t lean_float_beq(double, double); lean_object* l_Lean_Level_format(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_instToJsonMessageSeverity; lean_object* l_String_toName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_MessageLog_getInfoMessages___spec__3(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__22; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MessageLog_getInfoMessages___spec__5(lean_object*, size_t, size_t, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__31; -static lean_object* l_Lean_KernelException_toMessageData___closed__18; lean_object* l_Except_orElseLazy___rarg(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_instCoeArrayExpr___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageData_formatAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedMessageLog___closed__2; LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial(lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__49; static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__18; static lean_object* l_Lean_MessageLog_empty___closed__2; LEAN_EXPORT lean_object* l_Lean_MessageData_ofSyntax___elambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__13; -static lean_object* l_Lean_KernelException_toMessageData___closed__13; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__47; static lean_object* l_Lean_termM_x21_____closed__9; static lean_object* l___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____rarg___closed__30; static lean_object* l_Lean_instFromJsonSerialMessage___closed__1; @@ -505,7 +511,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_fromJsonMessageSeverit LEAN_EXPORT lean_object* l_Lean_instToMessageDataMessageData; LEAN_EXPORT lean_object* l_Lean_MessageSeverity_noConfusion___rarg___lambda__1(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_KernelException_mkCtx(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_toJsonSerialMessage____x40_Lean_Message___hyg_3221____closed__1; static lean_object* l_Lean_MessageData_paren___closed__1; LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax(lean_object*); @@ -517,26 +522,21 @@ static lean_object* l_Lean_MessageData_ofList___closed__1; static lean_object* l_Lean_instInhabitedMessageLog___closed__4; static lean_object* l_Lean_instImpl____x40_Lean_Message___hyg_606____closed__1; LEAN_EXPORT lean_object* l_Lean_MessageLog_hasUnreported___boxed(lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__26; lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); static lean_object* l_Lean_mkErrorStringWithPos___closed__2; LEAN_EXPORT lean_object* l_Lean_MessageData_formatAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_hasErrors___boxed(lean_object*); lean_object* l_Function_comp___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__47; LEAN_EXPORT lean_object* l_Lean_instImpl____x40_Lean_Message___hyg_606_; lean_object* l_Lean_PersistentArray_forM___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MessageData_format___closed__1; size_t lean_usize_add(size_t, size_t); -static lean_object* l_Lean_KernelException_toMessageData___closed__49; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_stringToMessageData___spec__3___at_Lean_stringToMessageData___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeMVarId(lean_object*); static lean_object* l_Lean_SerialMessage_toString___lambda__3___closed__1; static lean_object* l_Lean_instToMessageDataOption___rarg___closed__1; -static lean_object* l_Lean_KernelException_toMessageData___closed__14; static lean_object* l_Lean_termM_x21_____closed__10; -static lean_object* l_Lean_KernelException_toMessageData___closed__5; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_MessageData_ofList___closed__3; size_t lean_array_size(lean_object*); @@ -545,21 +545,18 @@ static lean_object* l_Lean_MessageData_instCoeArrayExpr___closed__3; LEAN_EXPORT lean_object* l_Lean_MessageLog_forM(lean_object*); uint8_t l_Lean_PersistentArray_isEmpty___rarg(lean_object*); static lean_object* l_Lean_MessageData_ofList___closed__8; -static lean_object* l_Lean_KernelException_toMessageData___closed__11; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__20; -static lean_object* l_Lean_KernelException_toMessageData___closed__17; static lean_object* l_Lean_instToMessageDataString___closed__1; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_stringToMessageData___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_107_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_MessageLog_msgs___boxed(lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__21; LEAN_EXPORT lean_object* l_Lean_SerialMessage_toString___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____closed__16; size_t lean_usize_shift_left(size_t, size_t); static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_MessageLog_hasErrors___spec__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageSeverity_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__20; LEAN_EXPORT lean_object* l_Lean_instInhabitedBaseMessage(lean_object*); lean_object* l_Lean_ppGoal(lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -572,7 +569,9 @@ static lean_object* l_Lean_termM_x21_____closed__3; LEAN_EXPORT lean_object* l_Lean_instToMessageDataName; static lean_object* l_Lean_instToJsonMessageSeverity___closed__1; lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__21; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__3(lean_object*, lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__5; LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeExpr; LEAN_EXPORT lean_object* l_Lean_instToMessageDataSubarray___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Message_toString___boxed(lean_object*, lean_object*, lean_object*); @@ -583,7 +582,6 @@ static lean_object* l___private_Lean_Message_0__Lean_toJsonBaseMessage____x40_Le uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageLog_errorsToWarnings___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_kind___boxed(lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__23; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2877____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageData_formatAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToJsonBaseMessage(lean_object*); @@ -609,10 +607,12 @@ static lean_object* l_Lean_instToMessageDataOptionExpr___closed__2; static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__7; LEAN_EXPORT lean_object* l_Lean_MessageData_lazy(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_arrayExpr_toMessageData___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); -static lean_object* l_Lean_KernelException_toMessageData___closed__51; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__18; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__10; static lean_object* l_Lean_MessageData_andList___closed__1; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__37; lean_object* l_String_toSubstring_x27(lean_object*); +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__48; lean_object* l_Lean_RBTree_union___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_markAllReported(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageLog_errorsToWarnings___spec__3(size_t, size_t, lean_object*); @@ -624,7 +624,7 @@ 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_Lean_MessageData_ofConstName___elambda__1___boxed(lean_object*); static lean_object* l_Lean_MessageData_paren___closed__2; -static lean_object* l_Lean_KernelException_toMessageData___closed__52; +static lean_object* l_Lean_Kernel_Exception_toMessageData___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MessageLog_errorsToWarnings___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_format(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_MessageData_orList___spec__1(lean_object*, lean_object*); @@ -633,11 +633,11 @@ LEAN_EXPORT lean_object* l_Lean_MessageLog_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MessageLog_hasErrors___spec__4___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_MessageData_mkPPContext___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_lazy___elambda__1___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonSerialMessage____x40_Lean_Message___hyg_3357____spec__1___closed__1; static lean_object* l_Lean_MessageData_instCoeString___closed__2; static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__4; -static lean_object* l_Lean_KernelException_toMessageData___closed__27; LEAN_EXPORT lean_object* l_Lean_stringToMessageData___boxed(lean_object*); double lean_float_sub(double, double); lean_object* l___private_Init_Dynamic_0__Dynamic_get_x3fImpl___rarg(lean_object*, lean_object*); @@ -11600,7 +11600,7 @@ x_5 = l_Lean_indentD(x_4); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_KernelException_mkCtx(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; @@ -11616,7 +11616,7 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__1() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__1() { _start: { lean_object* x_1; @@ -11624,16 +11624,16 @@ x_1 = lean_mk_string_unchecked("(kernel) unknown constant '", 27, 27); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__2() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__1; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__3() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -11642,7 +11642,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__4() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__4() { _start: { lean_object* x_1; @@ -11650,16 +11650,16 @@ x_1 = lean_mk_string_unchecked("(kernel) constant has already been declared '", return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__5() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__4; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__6() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__6() { _start: { lean_object* x_1; @@ -11667,26 +11667,26 @@ x_1 = lean_mk_string_unchecked("(kernel) declaration type mismatch", 34, 34); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__7() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__6; +x_1 = l_Lean_Kernel_Exception_toMessageData___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_Lean_KernelException_toMessageData___closed__8() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__7; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__7; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__9() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__9() { _start: { lean_object* x_1; @@ -11694,16 +11694,16 @@ x_1 = lean_mk_string_unchecked("(kernel) declaration type mismatch, '", 37, 37); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__10() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__9; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__9; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__11() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__11() { _start: { lean_object* x_1; @@ -11711,16 +11711,16 @@ x_1 = lean_mk_string_unchecked("' has type", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__12() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__11; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__11; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__13() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__13() { _start: { lean_object* x_1; @@ -11728,16 +11728,16 @@ x_1 = lean_mk_string_unchecked("\nbut it is expected to have type", 32, 32); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__14() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__13; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__13; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__15() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -11746,7 +11746,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__16() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__16() { _start: { lean_object* x_1; @@ -11754,16 +11754,16 @@ x_1 = lean_mk_string_unchecked("(kernel) declaration has metavariables '", 40, 4 return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__17() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__16; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__16; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__18() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__18() { _start: { lean_object* x_1; @@ -11771,16 +11771,16 @@ x_1 = lean_mk_string_unchecked("(kernel) declaration has free variables '", 41, return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__19() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__18; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__18; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__20() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__20() { _start: { lean_object* x_1; @@ -11788,16 +11788,16 @@ x_1 = lean_mk_string_unchecked("(kernel) function expected", 26, 26); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__21() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__20; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__20; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__22() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__22() { _start: { lean_object* x_1; @@ -11805,16 +11805,16 @@ x_1 = lean_mk_string_unchecked("(kernel) type expected", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__23() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__22; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__22; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__24() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__24() { _start: { lean_object* x_1; @@ -11822,16 +11822,16 @@ x_1 = lean_mk_string_unchecked("(kernel) let-declaration type mismatch '", 40, 4 return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__25() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__25() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__24; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__24; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__26() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__26() { _start: { lean_object* x_1; @@ -11839,16 +11839,16 @@ x_1 = lean_mk_string_unchecked("(kernel) type mismatch at", 25, 25); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__27() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__26; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__26; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__28() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__28() { _start: { lean_object* x_1; @@ -11856,16 +11856,16 @@ x_1 = lean_mk_string_unchecked("application type mismatch", 25, 25); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__29() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__29() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__28; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__28; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__30() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__30() { _start: { lean_object* x_1; @@ -11873,16 +11873,16 @@ x_1 = lean_mk_string_unchecked("\nargument has type", 18, 18); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__31() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__31() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__30; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__30; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__32() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__32() { _start: { lean_object* x_1; @@ -11890,16 +11890,16 @@ x_1 = lean_mk_string_unchecked("\nbut function has type", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__33() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__32; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__32; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__34() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__34() { _start: { lean_object* x_1; @@ -11907,16 +11907,16 @@ x_1 = lean_mk_string_unchecked("(kernel) invalid projection", 27, 27); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__35() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__35() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__34; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__34; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__36() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__36() { _start: { lean_object* x_1; @@ -11924,16 +11924,16 @@ x_1 = lean_mk_string_unchecked("(kernel) type of theorem '", 26, 26); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__37() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__37() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__36; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__36; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__38() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__38() { _start: { lean_object* x_1; @@ -11941,16 +11941,16 @@ x_1 = lean_mk_string_unchecked("' is not a proposition", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__39() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__38; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__38; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__40() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__40() { _start: { lean_object* x_1; @@ -11958,16 +11958,16 @@ x_1 = lean_mk_string_unchecked("(kernel) ", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__41() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__41() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__40; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__40; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__42() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__42() { _start: { lean_object* x_1; @@ -11975,26 +11975,26 @@ x_1 = lean_mk_string_unchecked("(kernel) deterministic timeout", 30, 30); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__43() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__42; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__42; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__44() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__44() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__43; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__43; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__45() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__45() { _start: { lean_object* x_1; @@ -12002,26 +12002,26 @@ x_1 = lean_mk_string_unchecked("(kernel) excessive memory consumption detected", return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__46() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__45; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__45; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__47() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__47() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__46; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__46; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__48() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__48() { _start: { lean_object* x_1; @@ -12029,26 +12029,26 @@ x_1 = lean_mk_string_unchecked("(kernel) deep recursion detected", 32, 32); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__49() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__49() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__48; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__48; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__50() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__50() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__49; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__49; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__51() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__51() { _start: { lean_object* x_1; @@ -12056,26 +12056,26 @@ x_1 = lean_mk_string_unchecked("(kernel) interrupted", 20, 20); return x_1; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__52() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__51; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__51; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_KernelException_toMessageData___closed__53() { +static lean_object* _init_l_Lean_Kernel_Exception_toMessageData___closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_KernelException_toMessageData___closed__52; +x_1 = l_Lean_Kernel_Exception_toMessageData___closed__52; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_KernelException_toMessageData(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object* x_1, lean_object* x_2) { _start: { switch (lean_obj_tag(x_1)) { @@ -12089,16 +12089,16 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_1, 0); x_5 = lean_ctor_get(x_1, 1); x_6 = l_Lean_MessageData_ofName(x_5); -x_7 = l_Lean_KernelException_toMessageData___closed__2; +x_7 = l_Lean_Kernel_Exception_toMessageData___closed__2; lean_ctor_set_tag(x_1, 7); lean_ctor_set(x_1, 1, x_6); lean_ctor_set(x_1, 0, x_7); -x_8 = l_Lean_KernelException_toMessageData___closed__3; +x_8 = l_Lean_Kernel_Exception_toMessageData___closed__3; x_9 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_9, 0, x_1); lean_ctor_set(x_9, 1, x_8); x_10 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_11 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_4, x_10, x_2, x_9); +x_11 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_4, x_10, x_2, x_9); return x_11; } else @@ -12110,16 +12110,16 @@ lean_inc(x_13); lean_inc(x_12); lean_dec(x_1); x_14 = l_Lean_MessageData_ofName(x_13); -x_15 = l_Lean_KernelException_toMessageData___closed__2; +x_15 = l_Lean_Kernel_Exception_toMessageData___closed__2; x_16 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); -x_17 = l_Lean_KernelException_toMessageData___closed__3; +x_17 = l_Lean_Kernel_Exception_toMessageData___closed__3; 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_addMessageContextPartial___rarg___lambda__1___closed__1; -x_20 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_12, x_19, x_2, x_18); +x_20 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_12, x_19, x_2, x_18); return x_20; } } @@ -12134,16 +12134,16 @@ x_22 = lean_ctor_get(x_1, 0); x_23 = lean_ctor_get(x_1, 1); x_24 = 1; x_25 = l_Lean_MessageData_ofConstName(x_23, x_24); -x_26 = l_Lean_KernelException_toMessageData___closed__5; +x_26 = l_Lean_Kernel_Exception_toMessageData___closed__5; lean_ctor_set_tag(x_1, 7); lean_ctor_set(x_1, 1, x_25); lean_ctor_set(x_1, 0, x_26); -x_27 = l_Lean_KernelException_toMessageData___closed__3; +x_27 = l_Lean_Kernel_Exception_toMessageData___closed__3; x_28 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_28, 0, x_1); lean_ctor_set(x_28, 1, x_27); x_29 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_30 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_22, x_29, x_2, x_28); +x_30 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_22, x_29, x_2, x_28); return x_30; } else @@ -12156,16 +12156,16 @@ lean_inc(x_31); lean_dec(x_1); x_33 = 1; x_34 = l_Lean_MessageData_ofConstName(x_32, x_33); -x_35 = l_Lean_KernelException_toMessageData___closed__5; +x_35 = l_Lean_Kernel_Exception_toMessageData___closed__5; x_36 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_34); -x_37 = l_Lean_KernelException_toMessageData___closed__3; +x_37 = l_Lean_Kernel_Exception_toMessageData___closed__3; x_38 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_37); x_39 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_40 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_31, x_39, x_2, x_38); +x_40 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_31, x_39, x_2, x_38); return x_40; } } @@ -12195,11 +12195,11 @@ x_47 = lean_ctor_get(x_43, 2); lean_inc(x_47); lean_dec(x_43); x_48 = l_Lean_MessageData_ofName(x_46); -x_49 = l_Lean_KernelException_toMessageData___closed__10; +x_49 = l_Lean_Kernel_Exception_toMessageData___closed__10; x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_48); -x_51 = l_Lean_KernelException_toMessageData___closed__12; +x_51 = l_Lean_Kernel_Exception_toMessageData___closed__12; x_52 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); @@ -12207,7 +12207,7 @@ x_53 = l_Lean_indentExpr(x_45); x_54 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_54, 0, x_52); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_KernelException_toMessageData___closed__14; +x_55 = l_Lean_Kernel_Exception_toMessageData___closed__14; x_56 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_56, 0, x_54); lean_ctor_set(x_56, 1, x_55); @@ -12215,12 +12215,12 @@ x_57 = l_Lean_indentExpr(x_47); x_58 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_58, 0, x_56); lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_KernelException_toMessageData___closed__15; +x_59 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_60 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_60, 0, x_58); lean_ctor_set(x_60, 1, x_59); x_61 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_62 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_44, x_61, x_2, x_60); +x_62 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_44, x_61, x_2, x_60); return x_62; } case 2: @@ -12243,11 +12243,11 @@ x_68 = lean_ctor_get(x_64, 2); lean_inc(x_68); lean_dec(x_64); x_69 = l_Lean_MessageData_ofName(x_67); -x_70 = l_Lean_KernelException_toMessageData___closed__10; +x_70 = l_Lean_Kernel_Exception_toMessageData___closed__10; x_71 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_71, 0, x_70); lean_ctor_set(x_71, 1, x_69); -x_72 = l_Lean_KernelException_toMessageData___closed__12; +x_72 = l_Lean_Kernel_Exception_toMessageData___closed__12; x_73 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_73, 0, x_71); lean_ctor_set(x_73, 1, x_72); @@ -12255,7 +12255,7 @@ x_74 = l_Lean_indentExpr(x_66); x_75 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_75, 0, x_73); lean_ctor_set(x_75, 1, x_74); -x_76 = l_Lean_KernelException_toMessageData___closed__14; +x_76 = l_Lean_Kernel_Exception_toMessageData___closed__14; x_77 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_77, 0, x_75); lean_ctor_set(x_77, 1, x_76); @@ -12263,12 +12263,12 @@ x_78 = l_Lean_indentExpr(x_68); x_79 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_79, 0, x_77); lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean_KernelException_toMessageData___closed__15; +x_80 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_81 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_81, 0, x_79); lean_ctor_set(x_81, 1, x_80); x_82 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_83 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_65, x_82, x_2, x_81); +x_83 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_65, x_82, x_2, x_81); return x_83; } default: @@ -12279,8 +12279,8 @@ x_84 = lean_ctor_get(x_1, 0); lean_inc(x_84); lean_dec(x_1); x_85 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_86 = l_Lean_KernelException_toMessageData___closed__8; -x_87 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_84, x_85, x_2, x_86); +x_86 = l_Lean_Kernel_Exception_toMessageData___closed__8; +x_87 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_84, x_85, x_2, x_86); return x_87; } } @@ -12295,16 +12295,16 @@ lean_inc(x_89); lean_dec(x_1); x_90 = 1; x_91 = l_Lean_MessageData_ofConstName(x_89, x_90); -x_92 = l_Lean_KernelException_toMessageData___closed__17; +x_92 = l_Lean_Kernel_Exception_toMessageData___closed__17; x_93 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_93, 0, x_92); lean_ctor_set(x_93, 1, x_91); -x_94 = l_Lean_KernelException_toMessageData___closed__3; +x_94 = l_Lean_Kernel_Exception_toMessageData___closed__3; x_95 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_95, 0, x_93); lean_ctor_set(x_95, 1, x_94); x_96 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_97 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_88, x_96, x_2, x_95); +x_97 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_88, x_96, x_2, x_95); return x_97; } case 4: @@ -12317,16 +12317,16 @@ lean_inc(x_99); lean_dec(x_1); x_100 = 1; x_101 = l_Lean_MessageData_ofConstName(x_99, x_100); -x_102 = l_Lean_KernelException_toMessageData___closed__19; +x_102 = l_Lean_Kernel_Exception_toMessageData___closed__19; x_103 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_103, 0, x_102); lean_ctor_set(x_103, 1, x_101); -x_104 = l_Lean_KernelException_toMessageData___closed__3; +x_104 = l_Lean_Kernel_Exception_toMessageData___closed__3; x_105 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_105, 0, x_103); lean_ctor_set(x_105, 1, x_104); x_106 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_107 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_98, x_106, x_2, x_105); +x_107 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_98, x_106, x_2, x_105); return x_107; } case 5: @@ -12340,15 +12340,15 @@ x_110 = lean_ctor_get(x_1, 2); lean_inc(x_110); lean_dec(x_1); x_111 = l_Lean_indentExpr(x_110); -x_112 = l_Lean_KernelException_toMessageData___closed__21; +x_112 = l_Lean_Kernel_Exception_toMessageData___closed__21; x_113 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_113, 0, x_112); lean_ctor_set(x_113, 1, x_111); -x_114 = l_Lean_KernelException_toMessageData___closed__15; +x_114 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_115 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_115, 0, x_113); lean_ctor_set(x_115, 1, x_114); -x_116 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_108, x_109, x_2, x_115); +x_116 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_108, x_109, x_2, x_115); return x_116; } case 6: @@ -12362,15 +12362,15 @@ x_119 = lean_ctor_get(x_1, 2); lean_inc(x_119); lean_dec(x_1); x_120 = l_Lean_indentExpr(x_119); -x_121 = l_Lean_KernelException_toMessageData___closed__23; +x_121 = l_Lean_Kernel_Exception_toMessageData___closed__23; x_122 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_122, 0, x_121); lean_ctor_set(x_122, 1, x_120); -x_123 = l_Lean_KernelException_toMessageData___closed__15; +x_123 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_124 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_124, 0, x_122); lean_ctor_set(x_124, 1, x_123); -x_125 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_117, x_118, x_2, x_124); +x_125 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_117, x_118, x_2, x_124); return x_125; } case 7: @@ -12384,15 +12384,15 @@ x_128 = lean_ctor_get(x_1, 2); lean_inc(x_128); lean_dec(x_1); x_129 = l_Lean_MessageData_ofName(x_128); -x_130 = l_Lean_KernelException_toMessageData___closed__25; +x_130 = l_Lean_Kernel_Exception_toMessageData___closed__25; x_131 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_131, 0, x_130); lean_ctor_set(x_131, 1, x_129); -x_132 = l_Lean_KernelException_toMessageData___closed__3; +x_132 = l_Lean_Kernel_Exception_toMessageData___closed__3; x_133 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_133, 0, x_131); lean_ctor_set(x_133, 1, x_132); -x_134 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_126, x_127, x_2, x_133); +x_134 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_126, x_127, x_2, x_133); return x_134; } case 8: @@ -12406,15 +12406,15 @@ x_137 = lean_ctor_get(x_1, 2); lean_inc(x_137); lean_dec(x_1); x_138 = l_Lean_indentExpr(x_137); -x_139 = l_Lean_KernelException_toMessageData___closed__27; +x_139 = l_Lean_Kernel_Exception_toMessageData___closed__27; x_140 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_140, 0, x_139); lean_ctor_set(x_140, 1, x_138); -x_141 = l_Lean_KernelException_toMessageData___closed__15; +x_141 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_142 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_142, 0, x_140); lean_ctor_set(x_142, 1, x_141); -x_143 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_135, x_136, x_2, x_142); +x_143 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_135, x_136, x_2, x_142); return x_143; } case 9: @@ -12432,11 +12432,11 @@ x_148 = lean_ctor_get(x_1, 4); lean_inc(x_148); lean_dec(x_1); x_149 = l_Lean_indentExpr(x_146); -x_150 = l_Lean_KernelException_toMessageData___closed__29; +x_150 = l_Lean_Kernel_Exception_toMessageData___closed__29; x_151 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_151, 0, x_150); lean_ctor_set(x_151, 1, x_149); -x_152 = l_Lean_KernelException_toMessageData___closed__31; +x_152 = l_Lean_Kernel_Exception_toMessageData___closed__31; x_153 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_153, 0, x_151); lean_ctor_set(x_153, 1, x_152); @@ -12444,7 +12444,7 @@ x_154 = l_Lean_indentExpr(x_148); x_155 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_155, 0, x_153); lean_ctor_set(x_155, 1, x_154); -x_156 = l_Lean_KernelException_toMessageData___closed__33; +x_156 = l_Lean_Kernel_Exception_toMessageData___closed__33; x_157 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_157, 0, x_155); lean_ctor_set(x_157, 1, x_156); @@ -12452,11 +12452,11 @@ x_158 = l_Lean_indentExpr(x_147); x_159 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_159, 0, x_157); lean_ctor_set(x_159, 1, x_158); -x_160 = l_Lean_KernelException_toMessageData___closed__15; +x_160 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_161 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_161, 0, x_159); lean_ctor_set(x_161, 1, x_160); -x_162 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_144, x_145, x_2, x_161); +x_162 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_144, x_145, x_2, x_161); return x_162; } case 10: @@ -12470,15 +12470,15 @@ x_165 = lean_ctor_get(x_1, 2); lean_inc(x_165); lean_dec(x_1); x_166 = l_Lean_indentExpr(x_165); -x_167 = l_Lean_KernelException_toMessageData___closed__35; +x_167 = l_Lean_Kernel_Exception_toMessageData___closed__35; x_168 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_168, 0, x_167); lean_ctor_set(x_168, 1, x_166); -x_169 = l_Lean_KernelException_toMessageData___closed__15; +x_169 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_170 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_170, 0, x_168); lean_ctor_set(x_170, 1, x_169); -x_171 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_163, x_164, x_2, x_170); +x_171 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_163, x_164, x_2, x_170); return x_171; } case 11: @@ -12493,11 +12493,11 @@ lean_inc(x_174); lean_dec(x_1); x_175 = 1; x_176 = l_Lean_MessageData_ofConstName(x_173, x_175); -x_177 = l_Lean_KernelException_toMessageData___closed__37; +x_177 = l_Lean_Kernel_Exception_toMessageData___closed__37; x_178 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_178, 0, x_177); lean_ctor_set(x_178, 1, x_176); -x_179 = l_Lean_KernelException_toMessageData___closed__39; +x_179 = l_Lean_Kernel_Exception_toMessageData___closed__39; x_180 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_180, 0, x_178); lean_ctor_set(x_180, 1, x_179); @@ -12505,12 +12505,12 @@ x_181 = l_Lean_indentExpr(x_174); x_182 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_182, 0, x_180); lean_ctor_set(x_182, 1, x_181); -x_183 = l_Lean_KernelException_toMessageData___closed__15; +x_183 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_184 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_184, 0, x_182); lean_ctor_set(x_184, 1, x_183); x_185 = l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; -x_186 = l___private_Lean_Message_0__Lean_KernelException_mkCtx(x_172, x_185, x_2, x_184); +x_186 = l___private_Lean_Message_0__Lean_Kernel_Exception_mkCtx(x_172, x_185, x_2, x_184); return x_186; } case 12: @@ -12522,11 +12522,11 @@ lean_inc(x_187); lean_dec(x_1); x_188 = l_Lean_stringToMessageData(x_187); lean_dec(x_187); -x_189 = l_Lean_KernelException_toMessageData___closed__41; +x_189 = l_Lean_Kernel_Exception_toMessageData___closed__41; x_190 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_190, 0, x_189); lean_ctor_set(x_190, 1, x_188); -x_191 = l_Lean_KernelException_toMessageData___closed__15; +x_191 = l_Lean_Kernel_Exception_toMessageData___closed__15; x_192 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_192, 0, x_190); lean_ctor_set(x_192, 1, x_191); @@ -12536,28 +12536,28 @@ case 13: { lean_object* x_193; lean_dec(x_2); -x_193 = l_Lean_KernelException_toMessageData___closed__44; +x_193 = l_Lean_Kernel_Exception_toMessageData___closed__44; return x_193; } case 14: { lean_object* x_194; lean_dec(x_2); -x_194 = l_Lean_KernelException_toMessageData___closed__47; +x_194 = l_Lean_Kernel_Exception_toMessageData___closed__47; return x_194; } case 15: { lean_object* x_195; lean_dec(x_2); -x_195 = l_Lean_KernelException_toMessageData___closed__50; +x_195 = l_Lean_Kernel_Exception_toMessageData___closed__50; return x_195; } default: { lean_object* x_196; lean_dec(x_2); -x_196 = l_Lean_KernelException_toMessageData___closed__53; +x_196 = l_Lean_Kernel_Exception_toMessageData___closed__53; return x_196; } } @@ -13103,112 +13103,112 @@ l_Lean_toMessageList___closed__1 = _init_l_Lean_toMessageList___closed__1(); lean_mark_persistent(l_Lean_toMessageList___closed__1); l_Lean_toMessageList___closed__2 = _init_l_Lean_toMessageList___closed__2(); lean_mark_persistent(l_Lean_toMessageList___closed__2); -l_Lean_KernelException_toMessageData___closed__1 = _init_l_Lean_KernelException_toMessageData___closed__1(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__1); -l_Lean_KernelException_toMessageData___closed__2 = _init_l_Lean_KernelException_toMessageData___closed__2(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__2); -l_Lean_KernelException_toMessageData___closed__3 = _init_l_Lean_KernelException_toMessageData___closed__3(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__3); -l_Lean_KernelException_toMessageData___closed__4 = _init_l_Lean_KernelException_toMessageData___closed__4(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__4); -l_Lean_KernelException_toMessageData___closed__5 = _init_l_Lean_KernelException_toMessageData___closed__5(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__5); -l_Lean_KernelException_toMessageData___closed__6 = _init_l_Lean_KernelException_toMessageData___closed__6(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__6); -l_Lean_KernelException_toMessageData___closed__7 = _init_l_Lean_KernelException_toMessageData___closed__7(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__7); -l_Lean_KernelException_toMessageData___closed__8 = _init_l_Lean_KernelException_toMessageData___closed__8(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__8); -l_Lean_KernelException_toMessageData___closed__9 = _init_l_Lean_KernelException_toMessageData___closed__9(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__9); -l_Lean_KernelException_toMessageData___closed__10 = _init_l_Lean_KernelException_toMessageData___closed__10(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__10); -l_Lean_KernelException_toMessageData___closed__11 = _init_l_Lean_KernelException_toMessageData___closed__11(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__11); -l_Lean_KernelException_toMessageData___closed__12 = _init_l_Lean_KernelException_toMessageData___closed__12(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__12); -l_Lean_KernelException_toMessageData___closed__13 = _init_l_Lean_KernelException_toMessageData___closed__13(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__13); -l_Lean_KernelException_toMessageData___closed__14 = _init_l_Lean_KernelException_toMessageData___closed__14(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__14); -l_Lean_KernelException_toMessageData___closed__15 = _init_l_Lean_KernelException_toMessageData___closed__15(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__15); -l_Lean_KernelException_toMessageData___closed__16 = _init_l_Lean_KernelException_toMessageData___closed__16(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__16); -l_Lean_KernelException_toMessageData___closed__17 = _init_l_Lean_KernelException_toMessageData___closed__17(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__17); -l_Lean_KernelException_toMessageData___closed__18 = _init_l_Lean_KernelException_toMessageData___closed__18(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__18); -l_Lean_KernelException_toMessageData___closed__19 = _init_l_Lean_KernelException_toMessageData___closed__19(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__19); -l_Lean_KernelException_toMessageData___closed__20 = _init_l_Lean_KernelException_toMessageData___closed__20(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__20); -l_Lean_KernelException_toMessageData___closed__21 = _init_l_Lean_KernelException_toMessageData___closed__21(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__21); -l_Lean_KernelException_toMessageData___closed__22 = _init_l_Lean_KernelException_toMessageData___closed__22(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__22); -l_Lean_KernelException_toMessageData___closed__23 = _init_l_Lean_KernelException_toMessageData___closed__23(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__23); -l_Lean_KernelException_toMessageData___closed__24 = _init_l_Lean_KernelException_toMessageData___closed__24(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__24); -l_Lean_KernelException_toMessageData___closed__25 = _init_l_Lean_KernelException_toMessageData___closed__25(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__25); -l_Lean_KernelException_toMessageData___closed__26 = _init_l_Lean_KernelException_toMessageData___closed__26(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__26); -l_Lean_KernelException_toMessageData___closed__27 = _init_l_Lean_KernelException_toMessageData___closed__27(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__27); -l_Lean_KernelException_toMessageData___closed__28 = _init_l_Lean_KernelException_toMessageData___closed__28(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__28); -l_Lean_KernelException_toMessageData___closed__29 = _init_l_Lean_KernelException_toMessageData___closed__29(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__29); -l_Lean_KernelException_toMessageData___closed__30 = _init_l_Lean_KernelException_toMessageData___closed__30(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__30); -l_Lean_KernelException_toMessageData___closed__31 = _init_l_Lean_KernelException_toMessageData___closed__31(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__31); -l_Lean_KernelException_toMessageData___closed__32 = _init_l_Lean_KernelException_toMessageData___closed__32(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__32); -l_Lean_KernelException_toMessageData___closed__33 = _init_l_Lean_KernelException_toMessageData___closed__33(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__33); -l_Lean_KernelException_toMessageData___closed__34 = _init_l_Lean_KernelException_toMessageData___closed__34(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__34); -l_Lean_KernelException_toMessageData___closed__35 = _init_l_Lean_KernelException_toMessageData___closed__35(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__35); -l_Lean_KernelException_toMessageData___closed__36 = _init_l_Lean_KernelException_toMessageData___closed__36(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__36); -l_Lean_KernelException_toMessageData___closed__37 = _init_l_Lean_KernelException_toMessageData___closed__37(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__37); -l_Lean_KernelException_toMessageData___closed__38 = _init_l_Lean_KernelException_toMessageData___closed__38(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__38); -l_Lean_KernelException_toMessageData___closed__39 = _init_l_Lean_KernelException_toMessageData___closed__39(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__39); -l_Lean_KernelException_toMessageData___closed__40 = _init_l_Lean_KernelException_toMessageData___closed__40(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__40); -l_Lean_KernelException_toMessageData___closed__41 = _init_l_Lean_KernelException_toMessageData___closed__41(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__41); -l_Lean_KernelException_toMessageData___closed__42 = _init_l_Lean_KernelException_toMessageData___closed__42(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__42); -l_Lean_KernelException_toMessageData___closed__43 = _init_l_Lean_KernelException_toMessageData___closed__43(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__43); -l_Lean_KernelException_toMessageData___closed__44 = _init_l_Lean_KernelException_toMessageData___closed__44(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__44); -l_Lean_KernelException_toMessageData___closed__45 = _init_l_Lean_KernelException_toMessageData___closed__45(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__45); -l_Lean_KernelException_toMessageData___closed__46 = _init_l_Lean_KernelException_toMessageData___closed__46(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__46); -l_Lean_KernelException_toMessageData___closed__47 = _init_l_Lean_KernelException_toMessageData___closed__47(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__47); -l_Lean_KernelException_toMessageData___closed__48 = _init_l_Lean_KernelException_toMessageData___closed__48(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__48); -l_Lean_KernelException_toMessageData___closed__49 = _init_l_Lean_KernelException_toMessageData___closed__49(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__49); -l_Lean_KernelException_toMessageData___closed__50 = _init_l_Lean_KernelException_toMessageData___closed__50(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__50); -l_Lean_KernelException_toMessageData___closed__51 = _init_l_Lean_KernelException_toMessageData___closed__51(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__51); -l_Lean_KernelException_toMessageData___closed__52 = _init_l_Lean_KernelException_toMessageData___closed__52(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__52); -l_Lean_KernelException_toMessageData___closed__53 = _init_l_Lean_KernelException_toMessageData___closed__53(); -lean_mark_persistent(l_Lean_KernelException_toMessageData___closed__53); +l_Lean_Kernel_Exception_toMessageData___closed__1 = _init_l_Lean_Kernel_Exception_toMessageData___closed__1(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__1); +l_Lean_Kernel_Exception_toMessageData___closed__2 = _init_l_Lean_Kernel_Exception_toMessageData___closed__2(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__2); +l_Lean_Kernel_Exception_toMessageData___closed__3 = _init_l_Lean_Kernel_Exception_toMessageData___closed__3(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__3); +l_Lean_Kernel_Exception_toMessageData___closed__4 = _init_l_Lean_Kernel_Exception_toMessageData___closed__4(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__4); +l_Lean_Kernel_Exception_toMessageData___closed__5 = _init_l_Lean_Kernel_Exception_toMessageData___closed__5(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__5); +l_Lean_Kernel_Exception_toMessageData___closed__6 = _init_l_Lean_Kernel_Exception_toMessageData___closed__6(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__6); +l_Lean_Kernel_Exception_toMessageData___closed__7 = _init_l_Lean_Kernel_Exception_toMessageData___closed__7(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__7); +l_Lean_Kernel_Exception_toMessageData___closed__8 = _init_l_Lean_Kernel_Exception_toMessageData___closed__8(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__8); +l_Lean_Kernel_Exception_toMessageData___closed__9 = _init_l_Lean_Kernel_Exception_toMessageData___closed__9(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__9); +l_Lean_Kernel_Exception_toMessageData___closed__10 = _init_l_Lean_Kernel_Exception_toMessageData___closed__10(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__10); +l_Lean_Kernel_Exception_toMessageData___closed__11 = _init_l_Lean_Kernel_Exception_toMessageData___closed__11(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__11); +l_Lean_Kernel_Exception_toMessageData___closed__12 = _init_l_Lean_Kernel_Exception_toMessageData___closed__12(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__12); +l_Lean_Kernel_Exception_toMessageData___closed__13 = _init_l_Lean_Kernel_Exception_toMessageData___closed__13(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__13); +l_Lean_Kernel_Exception_toMessageData___closed__14 = _init_l_Lean_Kernel_Exception_toMessageData___closed__14(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__14); +l_Lean_Kernel_Exception_toMessageData___closed__15 = _init_l_Lean_Kernel_Exception_toMessageData___closed__15(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__15); +l_Lean_Kernel_Exception_toMessageData___closed__16 = _init_l_Lean_Kernel_Exception_toMessageData___closed__16(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__16); +l_Lean_Kernel_Exception_toMessageData___closed__17 = _init_l_Lean_Kernel_Exception_toMessageData___closed__17(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__17); +l_Lean_Kernel_Exception_toMessageData___closed__18 = _init_l_Lean_Kernel_Exception_toMessageData___closed__18(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__18); +l_Lean_Kernel_Exception_toMessageData___closed__19 = _init_l_Lean_Kernel_Exception_toMessageData___closed__19(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__19); +l_Lean_Kernel_Exception_toMessageData___closed__20 = _init_l_Lean_Kernel_Exception_toMessageData___closed__20(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__20); +l_Lean_Kernel_Exception_toMessageData___closed__21 = _init_l_Lean_Kernel_Exception_toMessageData___closed__21(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__21); +l_Lean_Kernel_Exception_toMessageData___closed__22 = _init_l_Lean_Kernel_Exception_toMessageData___closed__22(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__22); +l_Lean_Kernel_Exception_toMessageData___closed__23 = _init_l_Lean_Kernel_Exception_toMessageData___closed__23(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__23); +l_Lean_Kernel_Exception_toMessageData___closed__24 = _init_l_Lean_Kernel_Exception_toMessageData___closed__24(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__24); +l_Lean_Kernel_Exception_toMessageData___closed__25 = _init_l_Lean_Kernel_Exception_toMessageData___closed__25(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__25); +l_Lean_Kernel_Exception_toMessageData___closed__26 = _init_l_Lean_Kernel_Exception_toMessageData___closed__26(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__26); +l_Lean_Kernel_Exception_toMessageData___closed__27 = _init_l_Lean_Kernel_Exception_toMessageData___closed__27(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__27); +l_Lean_Kernel_Exception_toMessageData___closed__28 = _init_l_Lean_Kernel_Exception_toMessageData___closed__28(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__28); +l_Lean_Kernel_Exception_toMessageData___closed__29 = _init_l_Lean_Kernel_Exception_toMessageData___closed__29(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__29); +l_Lean_Kernel_Exception_toMessageData___closed__30 = _init_l_Lean_Kernel_Exception_toMessageData___closed__30(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__30); +l_Lean_Kernel_Exception_toMessageData___closed__31 = _init_l_Lean_Kernel_Exception_toMessageData___closed__31(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__31); +l_Lean_Kernel_Exception_toMessageData___closed__32 = _init_l_Lean_Kernel_Exception_toMessageData___closed__32(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__32); +l_Lean_Kernel_Exception_toMessageData___closed__33 = _init_l_Lean_Kernel_Exception_toMessageData___closed__33(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__33); +l_Lean_Kernel_Exception_toMessageData___closed__34 = _init_l_Lean_Kernel_Exception_toMessageData___closed__34(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__34); +l_Lean_Kernel_Exception_toMessageData___closed__35 = _init_l_Lean_Kernel_Exception_toMessageData___closed__35(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__35); +l_Lean_Kernel_Exception_toMessageData___closed__36 = _init_l_Lean_Kernel_Exception_toMessageData___closed__36(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__36); +l_Lean_Kernel_Exception_toMessageData___closed__37 = _init_l_Lean_Kernel_Exception_toMessageData___closed__37(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__37); +l_Lean_Kernel_Exception_toMessageData___closed__38 = _init_l_Lean_Kernel_Exception_toMessageData___closed__38(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__38); +l_Lean_Kernel_Exception_toMessageData___closed__39 = _init_l_Lean_Kernel_Exception_toMessageData___closed__39(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__39); +l_Lean_Kernel_Exception_toMessageData___closed__40 = _init_l_Lean_Kernel_Exception_toMessageData___closed__40(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__40); +l_Lean_Kernel_Exception_toMessageData___closed__41 = _init_l_Lean_Kernel_Exception_toMessageData___closed__41(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__41); +l_Lean_Kernel_Exception_toMessageData___closed__42 = _init_l_Lean_Kernel_Exception_toMessageData___closed__42(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__42); +l_Lean_Kernel_Exception_toMessageData___closed__43 = _init_l_Lean_Kernel_Exception_toMessageData___closed__43(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__43); +l_Lean_Kernel_Exception_toMessageData___closed__44 = _init_l_Lean_Kernel_Exception_toMessageData___closed__44(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__44); +l_Lean_Kernel_Exception_toMessageData___closed__45 = _init_l_Lean_Kernel_Exception_toMessageData___closed__45(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__45); +l_Lean_Kernel_Exception_toMessageData___closed__46 = _init_l_Lean_Kernel_Exception_toMessageData___closed__46(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__46); +l_Lean_Kernel_Exception_toMessageData___closed__47 = _init_l_Lean_Kernel_Exception_toMessageData___closed__47(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__47); +l_Lean_Kernel_Exception_toMessageData___closed__48 = _init_l_Lean_Kernel_Exception_toMessageData___closed__48(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__48); +l_Lean_Kernel_Exception_toMessageData___closed__49 = _init_l_Lean_Kernel_Exception_toMessageData___closed__49(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__49); +l_Lean_Kernel_Exception_toMessageData___closed__50 = _init_l_Lean_Kernel_Exception_toMessageData___closed__50(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__50); +l_Lean_Kernel_Exception_toMessageData___closed__51 = _init_l_Lean_Kernel_Exception_toMessageData___closed__51(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__51); +l_Lean_Kernel_Exception_toMessageData___closed__52 = _init_l_Lean_Kernel_Exception_toMessageData___closed__52(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__52); +l_Lean_Kernel_Exception_toMessageData___closed__53 = _init_l_Lean_Kernel_Exception_toMessageData___closed__53(); +lean_mark_persistent(l_Lean_Kernel_Exception_toMessageData___closed__53); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/AppBuilder.c b/stage0/stdlib/Lean/Meta/AppBuilder.c index dbc16ff6c192..aa6b28d1562e 100644 --- a/stage0/stdlib/Lean/Meta/AppBuilder.c +++ b/stage0/stdlib/Lean/Meta/AppBuilder.c @@ -126,7 +126,7 @@ static lean_object* l_Lean_Meta_congrArg_x3f___closed__1; static lean_object* l_Lean_Meta_congrArg_x3f___lambda__4___closed__9; static lean_object* l_Lean_Meta_mkImpCongr___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___at_Lean_Meta_mkAppOptM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppOptMAux___closed__4; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_hasAssignableMVar___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMFinal___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -13103,7 +13103,8 @@ x_36 = lean_ctor_get(x_33, 1); x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_environment_find(x_37, x_31); +x_38 = l_Lean_Environment_find_x3f(x_37, x_31); +lean_dec(x_31); if (lean_obj_tag(x_38) == 0) { lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; @@ -13308,7 +13309,8 @@ lean_dec(x_33); x_101 = lean_ctor_get(x_99, 0); lean_inc(x_101); lean_dec(x_99); -x_102 = lean_environment_find(x_101, x_31); +x_102 = l_Lean_Environment_find_x3f(x_101, x_31); +lean_dec(x_31); if (lean_obj_tag(x_102) == 0) { lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; diff --git a/stage0/stdlib/Lean/Meta/Basic.c b/stage0/stdlib/Lean/Meta/Basic.c index 10672b3302c2..cd492f024e6f 100644 --- a/stage0/stdlib/Lean/Meta/Basic.c +++ b/stage0/stdlib/Lean/Meta/Basic.c @@ -295,6 +295,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_FVarId_throwUnknown___spec_ LEAN_EXPORT uint8_t l_Lean_Meta_instInhabitedProjReductionKind; LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_FVarId_hasForwardDeps___spec__49(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_isInductivePredicate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2486____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_withInstImplicitAsImplict(lean_object*); @@ -304,7 +305,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_isLevelDefEq(lean_object*, lean_object*, le LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_FVarId_hasForwardDeps___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaMetaTelescope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_Meta_instInhabitedPostponedEntry___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_FVarId_hasForwardDeps___spec__39___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -404,7 +405,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Basic_ LEAN_EXPORT lean_object* l_Lean_getReducibilityStatus___at___private_Lean_Meta_Basic_0__Lean_Meta_getDefInfoTemp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isInductivePredicate___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyInferTypeCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_FVarId_hasForwardDeps___spec__35___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_findDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -415,14 +415,12 @@ static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_reprConfig____x40_L LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_toUInt64___boxed(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_FVarId_hasForwardDeps___spec__48(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_recordSynthPendingFailure___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaBoundedTelescope___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_reprConfig____x40_Lean_Meta_Basic___hyg_579____closed__49; LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_index(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_resetCache___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withIncSynthPending___rarg(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_MVarId_getDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -559,7 +557,6 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_numLevelParams(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapForallTelescope_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkLevelStuckErrorMessage(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_FVarId_hasForwardDeps___spec__43___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_recordUnfold___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -604,6 +601,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_ProjReductionKind_noConfusion___rarg___lamb lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_reprConfig____x40_Lean_Meta_Basic___hyg_579____closed__44; +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_recordSynthPendingFailure___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isReducible___at___private_Lean_Meta_Basic_0__Lean_Meta_getDefInfoTemp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -993,7 +991,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at___private_Lean_Meta_Basic_0 uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_sortFVarIds___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___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_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_FVarId_hasForwardDeps___spec__31(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshExprMVar(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_reprConfig____x40_Lean_Meta_Basic___hyg_579_(lean_object*, lean_object*); @@ -1070,6 +1068,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___rarg(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_19460____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_orelseMergeErrorsImp(lean_object*); +lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1078,6 +1077,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_resetCache___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_toUInt64___boxed(lean_object*); lean_object* l_Lean_Core_withRestoreOrSaveFull___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedAbstractMVarsResult___closed__1; +lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_reprProjReductionKind____x40_Lean_Meta_Basic___hyg_210____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ProjReductionKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -1285,6 +1285,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___rarg___lambda__1___boxed(le LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassImp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_reprProjReductionKind____x40_Lean_Meta_Basic___hyg_210____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_useEtaStruct___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_reprConfig____x40_Lean_Meta_Basic___hyg_579____closed__36; static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_withLocalDeclsD___spec__1___rarg(size_t, size_t, lean_object*); @@ -6975,7 +6976,7 @@ lean_dec(x_76); x_79 = lean_ctor_get(x_77, 0); lean_inc(x_79); lean_dec(x_77); -x_80 = l_Lean_Kernel_isDiagnosticsEnabled(x_79); +x_80 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_79); lean_dec(x_79); if (x_80 == 0) { @@ -7029,7 +7030,7 @@ lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean x_86 = lean_ctor_get(x_83, 0); x_87 = lean_ctor_get(x_83, 4); lean_dec(x_87); -x_88 = l_Lean_Kernel_enableDiag(x_86, x_75); +x_88 = l_Lean_Kernel_Environment_enableDiag(x_86, x_75); x_89 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; lean_ctor_set(x_83, 4, x_89); lean_ctor_set(x_83, 0, x_88); @@ -7144,7 +7145,7 @@ lean_inc(x_114); lean_inc(x_113); lean_inc(x_112); lean_dec(x_83); -x_119 = l_Lean_Kernel_enableDiag(x_112, x_75); +x_119 = l_Lean_Kernel_Environment_enableDiag(x_112, x_75); x_120 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; x_121 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_121, 0, x_119); @@ -7378,7 +7379,7 @@ lean_dec(x_181); x_184 = lean_ctor_get(x_182, 0); lean_inc(x_184); lean_dec(x_182); -x_185 = l_Lean_Kernel_isDiagnosticsEnabled(x_184); +x_185 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_184); lean_dec(x_184); if (x_185 == 0) { @@ -7453,7 +7454,7 @@ if (lean_is_exclusive(x_188)) { lean_dec_ref(x_188); x_197 = lean_box(0); } -x_198 = l_Lean_Kernel_enableDiag(x_190, x_180); +x_198 = l_Lean_Kernel_Environment_enableDiag(x_190, x_180); x_199 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; if (lean_is_scalar(x_197)) { x_200 = lean_alloc_ctor(0, 8, 0); @@ -9704,12 +9705,12 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; x_24 = lean_ctor_get(x_19, 0); lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_24, x_1); +x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_24, x_1); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; x_26 = lean_unsigned_to_nat(1u); -x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); +x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); lean_ctor_set(x_19, 0, x_27); x_28 = lean_st_ref_set(x_3, x_18, x_20); x_29 = !lean_is_exclusive(x_28); @@ -9744,7 +9745,7 @@ lean_dec(x_25); x_36 = lean_unsigned_to_nat(1u); x_37 = lean_nat_add(x_35, x_36); lean_dec(x_35); -x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); +x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); lean_ctor_set(x_19, 0, x_38); x_39 = lean_st_ref_set(x_3, x_18, x_20); x_40 = !lean_is_exclusive(x_39); @@ -9784,12 +9785,12 @@ lean_inc(x_47); lean_inc(x_46); lean_dec(x_19); lean_inc(x_46); -x_50 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_46, x_1); +x_50 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_46, x_1); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; x_51 = lean_unsigned_to_nat(1u); -x_52 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_46, x_1, x_51); +x_52 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_46, x_1, x_51); x_53 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_53, 0, x_52); lean_ctor_set(x_53, 1, x_47); @@ -9826,7 +9827,7 @@ lean_dec(x_50); x_60 = lean_unsigned_to_nat(1u); x_61 = lean_nat_add(x_59, x_60); lean_dec(x_59); -x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_46, x_1, x_61); +x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_46, x_1, x_61); x_63 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_63, 0, x_62); lean_ctor_set(x_63, 1, x_47); @@ -9887,12 +9888,12 @@ if (lean_is_exclusive(x_19)) { x_77 = lean_box(0); } lean_inc(x_73); -x_78 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_73, x_1); +x_78 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_73, x_1); if (lean_obj_tag(x_78) == 0) { lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; x_79 = lean_unsigned_to_nat(1u); -x_80 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_73, x_1, x_79); +x_80 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_73, x_1, x_79); if (lean_is_scalar(x_77)) { x_81 = lean_alloc_ctor(0, 4, 0); } else { @@ -9938,7 +9939,7 @@ lean_dec(x_78); x_89 = lean_unsigned_to_nat(1u); x_90 = lean_nat_add(x_88, x_89); lean_dec(x_88); -x_91 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_73, x_1, x_90); +x_91 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_73, x_1, x_90); if (lean_is_scalar(x_77)) { x_92 = lean_alloc_ctor(0, 4, 0); } else { @@ -10053,12 +10054,12 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; x_24 = lean_ctor_get(x_19, 1); lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_24, x_1); +x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_24, x_1); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; x_26 = lean_unsigned_to_nat(1u); -x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); +x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); lean_ctor_set(x_19, 1, x_27); x_28 = lean_st_ref_set(x_3, x_18, x_20); x_29 = !lean_is_exclusive(x_28); @@ -10093,7 +10094,7 @@ lean_dec(x_25); x_36 = lean_unsigned_to_nat(1u); x_37 = lean_nat_add(x_35, x_36); lean_dec(x_35); -x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); +x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); lean_ctor_set(x_19, 1, x_38); x_39 = lean_st_ref_set(x_3, x_18, x_20); x_40 = !lean_is_exclusive(x_39); @@ -10133,12 +10134,12 @@ lean_inc(x_47); lean_inc(x_46); lean_dec(x_19); lean_inc(x_47); -x_50 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_47, x_1); +x_50 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_47, x_1); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; x_51 = lean_unsigned_to_nat(1u); -x_52 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_47, x_1, x_51); +x_52 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_47, x_1, x_51); x_53 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_53, 0, x_46); lean_ctor_set(x_53, 1, x_52); @@ -10175,7 +10176,7 @@ lean_dec(x_50); x_60 = lean_unsigned_to_nat(1u); x_61 = lean_nat_add(x_59, x_60); lean_dec(x_59); -x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_47, x_1, x_61); +x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_47, x_1, x_61); x_63 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_63, 0, x_46); lean_ctor_set(x_63, 1, x_62); @@ -10236,12 +10237,12 @@ if (lean_is_exclusive(x_19)) { x_77 = lean_box(0); } lean_inc(x_74); -x_78 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_74, x_1); +x_78 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_74, x_1); if (lean_obj_tag(x_78) == 0) { lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; x_79 = lean_unsigned_to_nat(1u); -x_80 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_74, x_1, x_79); +x_80 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_74, x_1, x_79); if (lean_is_scalar(x_77)) { x_81 = lean_alloc_ctor(0, 4, 0); } else { @@ -10287,7 +10288,7 @@ lean_dec(x_78); x_89 = lean_unsigned_to_nat(1u); x_90 = lean_nat_add(x_88, x_89); lean_dec(x_88); -x_91 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_74, x_1, x_90); +x_91 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_74, x_1, x_90); if (lean_is_scalar(x_77)) { x_92 = lean_alloc_ctor(0, 4, 0); } else { @@ -10402,12 +10403,12 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; x_24 = lean_ctor_get(x_19, 2); lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_24, x_1); +x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_24, x_1); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; x_26 = lean_unsigned_to_nat(1u); -x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); +x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); lean_ctor_set(x_19, 2, x_27); x_28 = lean_st_ref_set(x_3, x_18, x_20); x_29 = !lean_is_exclusive(x_28); @@ -10442,7 +10443,7 @@ lean_dec(x_25); x_36 = lean_unsigned_to_nat(1u); x_37 = lean_nat_add(x_35, x_36); lean_dec(x_35); -x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); +x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); lean_ctor_set(x_19, 2, x_38); x_39 = lean_st_ref_set(x_3, x_18, x_20); x_40 = !lean_is_exclusive(x_39); @@ -10482,12 +10483,12 @@ lean_inc(x_47); lean_inc(x_46); lean_dec(x_19); lean_inc(x_48); -x_50 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_48, x_1); +x_50 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_48, x_1); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; x_51 = lean_unsigned_to_nat(1u); -x_52 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_48, x_1, x_51); +x_52 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_48, x_1, x_51); x_53 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_53, 0, x_46); lean_ctor_set(x_53, 1, x_47); @@ -10524,7 +10525,7 @@ lean_dec(x_50); x_60 = lean_unsigned_to_nat(1u); x_61 = lean_nat_add(x_59, x_60); lean_dec(x_59); -x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_48, x_1, x_61); +x_62 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_48, x_1, x_61); x_63 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_63, 0, x_46); lean_ctor_set(x_63, 1, x_47); @@ -10585,12 +10586,12 @@ if (lean_is_exclusive(x_19)) { x_77 = lean_box(0); } lean_inc(x_75); -x_78 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_75, x_1); +x_78 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_75, x_1); if (lean_obj_tag(x_78) == 0) { lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; x_79 = lean_unsigned_to_nat(1u); -x_80 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_75, x_1, x_79); +x_80 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_75, x_1, x_79); if (lean_is_scalar(x_77)) { x_81 = lean_alloc_ctor(0, 4, 0); } else { @@ -10636,7 +10637,7 @@ lean_dec(x_78); x_89 = lean_unsigned_to_nat(1u); x_90 = lean_nat_add(x_88, x_89); lean_dec(x_88); -x_91 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_75, x_1, x_90); +x_91 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_75, x_1, x_90); if (lean_is_scalar(x_77)) { x_92 = lean_alloc_ctor(0, 4, 0); } else { @@ -13867,8 +13868,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_1); -x_12 = lean_environment_find(x_11, x_1); +x_12 = l_Lean_Environment_find_x3f(x_11, x_1); if (lean_obj_tag(x_12) == 0) { 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; @@ -13908,8 +13908,7 @@ lean_dec(x_7); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -lean_inc(x_1); -x_24 = lean_environment_find(x_23, x_1); +x_24 = l_Lean_Environment_find_x3f(x_23, x_1); if (lean_obj_tag(x_24) == 0) { uint8_t 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; @@ -23346,7 +23345,8 @@ if (x_21 == 0) lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_20, 1); -x_24 = lean_environment_main_module(x_11); +x_24 = l_Lean_Environment_mainModule(x_11); +lean_dec(x_11); lean_inc(x_7); lean_ctor_set(x_20, 1, x_7); lean_ctor_set(x_20, 0, x_24); @@ -23737,7 +23737,8 @@ x_140 = lean_ctor_get(x_20, 1); lean_inc(x_140); lean_inc(x_139); lean_dec(x_20); -x_141 = lean_environment_main_module(x_11); +x_141 = l_Lean_Environment_mainModule(x_11); +lean_dec(x_11); lean_inc(x_7); x_142 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_142, 0, x_141); @@ -24086,7 +24087,8 @@ lean_inc(x_22); x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_12); +x_24 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_25 = lean_ctor_get(x_22, 1); lean_inc(x_25); lean_dec(x_22); @@ -24603,7 +24605,8 @@ lean_inc(x_22); x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_12); +x_24 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_25 = lean_ctor_get(x_22, 1); lean_inc(x_25); lean_dec(x_22); @@ -25064,7 +25067,8 @@ if (x_26 == 0) lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; x_27 = lean_ctor_get(x_25, 0); x_28 = lean_ctor_get(x_25, 1); -x_29 = lean_environment_main_module(x_16); +x_29 = l_Lean_Environment_mainModule(x_16); +lean_dec(x_16); lean_inc(x_12); lean_ctor_set(x_25, 1, x_12); lean_ctor_set(x_25, 0, x_29); @@ -25456,7 +25460,8 @@ x_146 = lean_ctor_get(x_25, 1); lean_inc(x_146); lean_inc(x_145); lean_dec(x_25); -x_147 = lean_environment_main_module(x_16); +x_147 = l_Lean_Environment_mainModule(x_16); +lean_dec(x_16); lean_inc(x_12); x_148 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_148, 0, x_147); @@ -25770,7 +25775,8 @@ 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; uint8_t x_34; lean_object* x_35; x_28 = lean_ctor_get(x_26, 0); x_29 = lean_ctor_get(x_26, 1); -x_30 = lean_environment_main_module(x_17); +x_30 = l_Lean_Environment_mainModule(x_17); +lean_dec(x_17); lean_inc(x_13); lean_ctor_set(x_26, 1, x_13); lean_ctor_set(x_26, 0, x_30); @@ -26162,7 +26168,8 @@ x_147 = lean_ctor_get(x_26, 1); lean_inc(x_147); lean_inc(x_146); lean_dec(x_26); -x_148 = lean_environment_main_module(x_17); +x_148 = l_Lean_Environment_mainModule(x_17); +lean_dec(x_17); lean_inc(x_13); x_149 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_149, 0, x_148); @@ -26597,7 +26604,8 @@ if (x_24 == 0) lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; x_25 = lean_ctor_get(x_23, 0); x_26 = lean_ctor_get(x_23, 1); -x_27 = lean_environment_main_module(x_14); +x_27 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); lean_inc(x_10); lean_ctor_set(x_23, 1, x_10); lean_ctor_set(x_23, 0, x_27); @@ -26989,7 +26997,8 @@ x_143 = lean_ctor_get(x_23, 1); lean_inc(x_143); lean_inc(x_142); lean_dec(x_23); -x_144 = lean_environment_main_module(x_14); +x_144 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); lean_inc(x_10); x_145 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_145, 0, x_144); @@ -32615,8 +32624,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_1); -x_12 = lean_environment_find(x_11, x_1); +x_12 = l_Lean_Environment_find_x3f(x_11, x_1); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; @@ -32701,8 +32709,7 @@ lean_dec(x_7); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; @@ -47177,7 +47184,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -x_12 = lean_environment_find(x_11, x_1); +x_12 = l_Lean_Environment_find_x3f(x_11, x_1); if (lean_obj_tag(x_12) == 0) { uint8_t x_13; lean_object* x_14; @@ -47240,7 +47247,7 @@ lean_dec(x_7); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -x_27 = lean_environment_find(x_26, x_1); +x_27 = l_Lean_Environment_find_x3f(x_26, x_1); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; @@ -47306,6 +47313,15 @@ lean_dec(x_1); return x_8; } } +LEAN_EXPORT lean_object* l_Lean_Meta_isInductivePredicate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_isInductivePredicate(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_1); +return x_7; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_isListLevelDefEqAux(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { diff --git a/stage0/stdlib/Lean/Meta/CompletionName.c b/stage0/stdlib/Lean/Meta/CompletionName.c index 028f455fa4c0..1934162301bc 100644 --- a/stage0/stdlib/Lean/Meta/CompletionName.c +++ b/stage0/stdlib/Lean/Meta/CompletionName.c @@ -190,7 +190,6 @@ x_6 = l_Lean_TagDeclarationExtension_isTagged(x_5, x_1, x_2); if (x_6 == 0) { uint8_t x_7; -lean_inc(x_2); lean_inc(x_1); x_7 = l_Lean_isRecCore(x_1, x_2); if (x_7 == 0) diff --git a/stage0/stdlib/Lean/Meta/CongrTheorems.c b/stage0/stdlib/Lean/Meta/CongrTheorems.c index d70016435952..a3fbb65f0de4 100644 --- a/stage0/stdlib/Lean/Meta/CongrTheorems.c +++ b/stage0/stdlib/Lean/Meta/CongrTheorems.c @@ -10500,6 +10500,7 @@ else { uint8_t x_9; x_9 = l_Lean_Environment_isSafeDefinition(x_1, x_3); +lean_dec(x_3); return x_9; } } @@ -10508,6 +10509,7 @@ else uint8_t x_10; lean_dec(x_4); x_10 = l_Lean_Environment_isSafeDefinition(x_1, x_3); +lean_dec(x_3); return x_10; } } @@ -12335,7 +12337,6 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_5); x_12 = l_Lean_Environment_isSafeDefinition(x_11, x_5); if (x_12 == 0) { @@ -12370,7 +12371,6 @@ lean_dec(x_7); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -lean_inc(x_5); x_20 = l_Lean_Environment_isSafeDefinition(x_19, x_5); if (x_20 == 0) { diff --git a/stage0/stdlib/Lean/Meta/Constructions/CasesOn.c b/stage0/stdlib/Lean/Meta/Constructions/CasesOn.c index 48d57035cf21..ab387f47a7fe 100644 --- a/stage0/stdlib/Lean/Meta/Constructions/CasesOn.c +++ b/stage0/stdlib/Lean/Meta/Constructions/CasesOn.c @@ -30,6 +30,7 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExceptKernelException___at_Lean_mkCasesOn___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_setReducibilityStatusCore(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_setReducibilityStatus___at_Lean_mkCasesOn___spec__4___closed__2; +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCasesOn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -37,7 +38,6 @@ lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_obj static lean_object* l_Lean_mkCasesOn___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_setReducibilityStatus___at_Lean_mkCasesOn___spec__4___closed__4; -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setReducibilityStatus___at_Lean_mkCasesOn___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_cases_on(lean_object*, lean_object*); @@ -95,7 +95,7 @@ LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_mkCasesOn___spec_ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_4, 2); lean_inc(x_7); -x_8 = l_Lean_KernelException_toMessageData(x_1, x_7); +x_8 = l_Lean_Kernel_Exception_toMessageData(x_1, x_7); x_9 = l_Lean_throwError___at_Lean_mkCasesOn___spec__3(x_8, x_2, x_3, x_4, x_5, x_6); lean_dec(x_4); return x_9; diff --git a/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c b/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c index 0cb74b4bb6da..49b3f9f2acf8 100644 --- a/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c +++ b/stage0/stdlib/Lean/Meta/Constructions/NoConfusion.c @@ -98,6 +98,7 @@ lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object static lean_object* l_Lean_mkNoConfusionEnum_mkNoConfusionType___lambda__1___closed__2; extern lean_object* l_Lean_Meta_completionBlackListExt; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkNoConfusionEnum_mkToCtorIdx(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExceptKernelException___at_Lean_mkNoConfusionCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); @@ -128,7 +129,6 @@ lean_object* l_Lean_addAndCompile(lean_object*, lean_object*, lean_object*, lean lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_mkNoConfusionEnum_mkToCtorIdx___closed__9; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkNoConfusionEnum_mkToCtorIdx___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); static lean_object* l_Lean_setReducibilityStatus___at_Lean_mkNoConfusionCore___spec__4___closed__4; static lean_object* l_Lean_mkNoConfusionEnum_mkToCtorIdx___closed__3; LEAN_EXPORT lean_object* l_Lean_setReducibilityStatus___at_Lean_mkNoConfusionCore___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -197,7 +197,7 @@ LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_mkNoConfusionCore lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_4, 2); lean_inc(x_7); -x_8 = l_Lean_KernelException_toMessageData(x_1, x_7); +x_8 = l_Lean_Kernel_Exception_toMessageData(x_1, x_7); x_9 = l_Lean_throwError___at_Lean_mkNoConfusionCore___spec__3(x_8, x_2, x_3, x_4, x_5, x_6); lean_dec(x_4); return x_9; diff --git a/stage0/stdlib/Lean/Meta/CtorRecognizer.c b/stage0/stdlib/Lean/Meta/CtorRecognizer.c index 0f34391a2664..a56d7a12e30e 100644 --- a/stage0/stdlib/Lean/Meta/CtorRecognizer.c +++ b/stage0/stdlib/Lean/Meta/CtorRecognizer.c @@ -26,13 +26,14 @@ lean_object* l_Lean_Meta_isOffset_x3f(lean_object*, lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_Meta_isConstructorApp_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isConstructorAppCore_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isConstructorApp_x27___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_constructorApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_constructorApp_x3f___closed__1; extern lean_object* l___private_Lean_Expr_0__Lean_natAddFn; lean_object* lean_st_ref_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getConstructorVal_x3f___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_Meta_isConstructorApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -54,7 +55,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getCo _start: { lean_object* x_3; -x_3 = lean_environment_find(x_1, x_2); +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_3) == 0) { lean_object* x_4; @@ -114,6 +115,15 @@ return x_12; } } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getConstructorVal_x3f___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getConstructorVal_x3f(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_isConstructorAppCore_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -135,6 +145,7 @@ x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); lean_dec(x_11); x_13 = l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getConstructorVal_x3f(x_12, x_8); +lean_dec(x_8); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; @@ -226,6 +237,7 @@ x_35 = lean_ctor_get(x_33, 0); lean_inc(x_35); lean_dec(x_33); x_36 = l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getConstructorVal_x3f(x_35, x_8); +lean_dec(x_8); if (lean_obj_tag(x_36) == 0) { lean_object* x_37; lean_object* x_38; @@ -1450,6 +1462,7 @@ x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); lean_dec(x_15); x_17 = l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getConstructorVal_x3f(x_16, x_12); +lean_dec(x_12); if (lean_obj_tag(x_17) == 0) { lean_object* x_18; @@ -1566,6 +1579,7 @@ x_51 = lean_ctor_get(x_49, 0); lean_inc(x_51); lean_dec(x_49); x_52 = l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getConstructorVal_x3f(x_51, x_12); +lean_dec(x_12); if (lean_obj_tag(x_52) == 0) { lean_object* x_53; lean_object* x_54; @@ -1683,6 +1697,7 @@ x_82 = lean_ctor_get(x_79, 0); lean_inc(x_82); lean_dec(x_79); x_83 = l___private_Lean_Meta_CtorRecognizer_0__Lean_Meta_getConstructorVal_x3f(x_82, x_77); +lean_dec(x_77); if (lean_obj_tag(x_83) == 0) { lean_object* x_84; lean_object* x_85; diff --git a/stage0/stdlib/Lean/Meta/Diagnostics.c b/stage0/stdlib/Lean/Meta/Diagnostics.c index b0c1abf983af..d3b32dce3df7 100644 --- a/stage0/stdlib/Lean/Meta/Diagnostics.c +++ b/stage0/stdlib/Lean/Meta/Diagnostics.c @@ -15,7 +15,6 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2___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_PersistentHashMap_Node_isEmpty___rarg(lean_object*); -extern lean_object* l_Lean_diagExt; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSynthPendingFailure___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSynthPendingFailure___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -73,7 +72,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedDiagSummary; LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_isInstanceCore(lean_object*, lean_object*); -lean_object* l_Lean_EnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSynthPendingFailure___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_EXPORT lean_object* l_Lean_Meta_reportDiag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); @@ -150,7 +148,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_mkDiagSy static lean_object* l_Lean_Meta_subCounters___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_appendSection___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Kernel_instInhabitedDiagnostics; extern lean_object* l_Lean_Meta_maxSynthPendingDepth; static lean_object* l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummary___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -211,7 +208,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary_ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_reportDiag___closed__16; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__1; lean_object* l_Lean_isDiagnosticsEnabled(lean_object*, lean_object*, lean_object*); @@ -5210,19 +5206,11 @@ static lean_object* _init_l_Lean_Meta_reportDiag___closed__5() { _start: { lean_object* x_1; -x_1 = l_Lean_diagExt; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_reportDiag___closed__6() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_unchecked("unfolded declarations", 21, 21); return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__7() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__6() { _start: { lean_object* x_1; @@ -5230,7 +5218,7 @@ x_1 = lean_mk_string_unchecked("unfolded instances", 18, 18); return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__8() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__7() { _start: { lean_object* x_1; @@ -5238,7 +5226,7 @@ x_1 = lean_mk_string_unchecked("unfolded reducible declarations", 31, 31); return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__9() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__8() { _start: { lean_object* x_1; @@ -5246,7 +5234,7 @@ x_1 = lean_mk_string_unchecked("used instances", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__10() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__9() { _start: { lean_object* x_1; @@ -5254,7 +5242,7 @@ x_1 = l_Lean_Meta_maxSynthPendingDepth; return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__11() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__10() { _start: { lean_object* x_1; @@ -5262,7 +5250,7 @@ x_1 = lean_mk_string_unchecked("max synth pending failures (maxSynthPendingDepth return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__12() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__11() { _start: { lean_object* x_1; @@ -5270,7 +5258,7 @@ x_1 = lean_mk_string_unchecked("), use `set_option maxSynthPendingDepth ` return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__13() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__12() { _start: { lean_object* x_1; @@ -5278,7 +5266,7 @@ x_1 = lean_mk_string_unchecked("heuristic for solving `f a =\?= f b`", 35, 35); return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__14() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__13() { _start: { lean_object* x_1; @@ -5286,21 +5274,21 @@ x_1 = lean_mk_string_unchecked("use `set_option diagnostics.threshold ` to return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__15() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_reportDiag___closed__14; +x_1 = l_Lean_Meta_reportDiag___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_Lean_Meta_reportDiag___closed__16() { +static lean_object* _init_l_Lean_Meta_reportDiag___closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_reportDiag___closed__15; +x_1 = l_Lean_Meta_reportDiag___closed__14; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } @@ -5454,542 +5442,541 @@ x_53 = lean_st_ref_get(x_4, x_52); x_54 = !lean_is_exclusive(x_53); if (x_54 == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; x_55 = lean_ctor_get(x_53, 0); x_56 = lean_ctor_get(x_53, 1); x_57 = lean_ctor_get(x_55, 0); lean_inc(x_57); lean_dec(x_55); -x_58 = l_Lean_Kernel_instInhabitedDiagnostics; -x_59 = l_Lean_Meta_reportDiag___closed__5; -x_60 = l_Lean_EnvExtension_getState___rarg(x_58, x_59, x_57); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); lean_dec(x_57); -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -lean_dec(x_60); -x_62 = l_Lean_Meta_reportDiag___closed__4; -x_63 = l_Lean_Meta_mkDiagSummary(x_62, x_61, x_38, x_1, x_2, x_3, x_4, x_56); -if (lean_obj_tag(x_63) == 0) -{ -uint8_t x_64; -x_64 = !lean_is_exclusive(x_63); -if (x_64 == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_65 = lean_ctor_get(x_63, 0); -x_66 = lean_ctor_get(x_63, 1); -x_67 = l_Lean_MessageData_nil; -x_68 = l_Lean_Meta_mkDiagSummaryForUnfolded___closed__2; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +lean_dec(x_58); +x_60 = l_Lean_Meta_reportDiag___closed__4; +x_61 = l_Lean_Meta_mkDiagSummary(x_60, x_59, x_38, x_1, x_2, x_3, x_4, x_56); +if (lean_obj_tag(x_61) == 0) +{ +uint8_t x_62; +x_62 = !lean_is_exclusive(x_61); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_63 = lean_ctor_get(x_61, 0); +x_64 = lean_ctor_get(x_61, 1); +x_65 = l_Lean_MessageData_nil; +x_66 = l_Lean_Meta_mkDiagSummaryForUnfolded___closed__2; +x_67 = l_Lean_Meta_reportDiag___closed__5; +x_68 = l_Lean_Meta_appendSection(x_65, x_66, x_67, x_23, x_25); x_69 = l_Lean_Meta_reportDiag___closed__6; -x_70 = l_Lean_Meta_appendSection(x_67, x_68, x_69, x_23, x_25); +x_70 = l_Lean_Meta_appendSection(x_68, x_66, x_69, x_27, x_25); x_71 = l_Lean_Meta_reportDiag___closed__7; -x_72 = l_Lean_Meta_appendSection(x_70, x_68, x_71, x_27, x_25); -x_73 = l_Lean_Meta_reportDiag___closed__8; -x_74 = l_Lean_Meta_appendSection(x_72, x_68, x_73, x_30, x_25); -x_75 = l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__2; -x_76 = l_Lean_Meta_reportDiag___closed__9; -x_77 = l_Lean_Meta_appendSection(x_74, x_75, x_76, x_43, x_25); -x_78 = lean_ctor_get(x_3, 2); -lean_inc(x_78); -x_79 = l_Lean_Meta_reportDiag___closed__10; -x_80 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_78, x_79); -lean_dec(x_78); -x_81 = l___private_Init_Data_Repr_0__Nat_reprFast(x_80); +x_72 = l_Lean_Meta_appendSection(x_70, x_66, x_71, x_30, x_25); +x_73 = l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__2; +x_74 = l_Lean_Meta_reportDiag___closed__8; +x_75 = l_Lean_Meta_appendSection(x_72, x_73, x_74, x_43, x_25); +x_76 = lean_ctor_get(x_3, 2); +lean_inc(x_76); +x_77 = l_Lean_Meta_reportDiag___closed__9; +x_78 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_76, x_77); +lean_dec(x_76); +x_79 = l___private_Init_Data_Repr_0__Nat_reprFast(x_78); +x_80 = l_Lean_Meta_reportDiag___closed__10; +x_81 = lean_string_append(x_80, x_79); +lean_dec(x_79); x_82 = l_Lean_Meta_reportDiag___closed__11; -x_83 = lean_string_append(x_82, x_81); -lean_dec(x_81); -x_84 = l_Lean_Meta_reportDiag___closed__12; -x_85 = lean_string_append(x_83, x_84); -x_86 = l_Lean_Meta_appendSection(x_77, x_75, x_85, x_51, x_21); -x_87 = l_Lean_Meta_reportDiag___closed__13; -x_88 = l_Lean_Meta_appendSection(x_86, x_37, x_87, x_40, x_25); -x_89 = l_Lean_Meta_appendSection(x_88, x_62, x_69, x_65, x_25); -if (lean_obj_tag(x_89) == 0) +x_83 = lean_string_append(x_81, x_82); +x_84 = l_Lean_Meta_appendSection(x_75, x_73, x_83, x_51, x_21); +x_85 = l_Lean_Meta_reportDiag___closed__12; +x_86 = l_Lean_Meta_appendSection(x_84, x_37, x_85, x_40, x_25); +x_87 = l_Lean_Meta_appendSection(x_86, x_60, x_67, x_63, x_25); +if (lean_obj_tag(x_87) == 0) { -lean_object* x_90; lean_object* x_91; +lean_object* x_88; lean_object* x_89; lean_free_object(x_53); -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_90, 0); -lean_inc(x_91); -switch (lean_obj_tag(x_91)) { +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +switch (lean_obj_tag(x_89)) { case 0: { -uint8_t x_92; -x_92 = !lean_is_exclusive(x_90); -if (x_92 == 0) +uint8_t x_90; +x_90 = !lean_is_exclusive(x_88); +if (x_90 == 0) { -lean_object* x_93; lean_object* x_94; -x_93 = lean_ctor_get(x_90, 1); -x_94 = lean_ctor_get(x_90, 0); -lean_dec(x_94); -if (lean_obj_tag(x_93) == 0) +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_88, 1); +x_92 = lean_ctor_get(x_88, 0); +lean_dec(x_92); +if (lean_obj_tag(x_91) == 0) { -lean_object* x_95; -lean_free_object(x_90); -lean_dec(x_89); +lean_object* x_93; +lean_free_object(x_88); +lean_dec(x_87); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_95 = lean_box(0); -lean_ctor_set(x_63, 0, x_95); -return x_63; +x_93 = lean_box(0); +lean_ctor_set(x_61, 0, x_93); +return x_61; } else { -lean_object* x_96; uint8_t x_97; lean_object* x_98; -lean_dec(x_93); -lean_free_object(x_63); -x_96 = l_Lean_Meta_reportDiag___closed__16; -lean_ctor_set_tag(x_90, 7); -lean_ctor_set(x_90, 1, x_96); -lean_ctor_set(x_90, 0, x_89); -x_97 = 0; -x_98 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_90, x_97, x_1, x_2, x_3, x_4, x_66); +lean_object* x_94; uint8_t x_95; lean_object* x_96; +lean_dec(x_91); +lean_free_object(x_61); +x_94 = l_Lean_Meta_reportDiag___closed__15; +lean_ctor_set_tag(x_88, 7); +lean_ctor_set(x_88, 1, x_94); +lean_ctor_set(x_88, 0, x_87); +x_95 = 0; +x_96 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_88, x_95, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_98; +return x_96; } } else { -lean_object* x_99; -x_99 = lean_ctor_get(x_90, 1); -lean_inc(x_99); -lean_dec(x_90); -if (lean_obj_tag(x_99) == 0) +lean_object* x_97; +x_97 = lean_ctor_get(x_88, 1); +lean_inc(x_97); +lean_dec(x_88); +if (lean_obj_tag(x_97) == 0) { -lean_object* x_100; -lean_dec(x_89); +lean_object* x_98; +lean_dec(x_87); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_100 = lean_box(0); -lean_ctor_set(x_63, 0, x_100); -return x_63; +x_98 = lean_box(0); +lean_ctor_set(x_61, 0, x_98); +return x_61; } else { -lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_object* x_104; -lean_dec(x_99); -lean_free_object(x_63); -x_101 = l_Lean_Meta_reportDiag___closed__16; -x_102 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_102, 0, x_89); -lean_ctor_set(x_102, 1, x_101); -x_103 = 0; -x_104 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_102, x_103, x_1, x_2, x_3, x_4, x_66); +lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; +lean_dec(x_97); +lean_free_object(x_61); +x_99 = l_Lean_Meta_reportDiag___closed__15; +x_100 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_100, 0, x_87); +lean_ctor_set(x_100, 1, x_99); +x_101 = 0; +x_102 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_100, x_101, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_104; +return x_102; } } } case 4: { -uint8_t x_105; -lean_dec(x_90); -lean_free_object(x_63); -x_105 = !lean_is_exclusive(x_91); -if (x_105 == 0) -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; -x_106 = lean_ctor_get(x_91, 1); -lean_dec(x_106); -x_107 = lean_ctor_get(x_91, 0); -lean_dec(x_107); -x_108 = l_Lean_Meta_reportDiag___closed__16; -lean_ctor_set_tag(x_91, 7); -lean_ctor_set(x_91, 1, x_108); -lean_ctor_set(x_91, 0, x_89); -x_109 = 0; -x_110 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_91, x_109, x_1, x_2, x_3, x_4, x_66); +uint8_t x_103; +lean_dec(x_88); +lean_free_object(x_61); +x_103 = !lean_is_exclusive(x_89); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; +x_104 = lean_ctor_get(x_89, 1); +lean_dec(x_104); +x_105 = lean_ctor_get(x_89, 0); +lean_dec(x_105); +x_106 = l_Lean_Meta_reportDiag___closed__15; +lean_ctor_set_tag(x_89, 7); +lean_ctor_set(x_89, 1, x_106); +lean_ctor_set(x_89, 0, x_87); +x_107 = 0; +x_108 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_89, x_107, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_110; +return x_108; } else { -lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; -lean_dec(x_91); -x_111 = l_Lean_Meta_reportDiag___closed__16; -x_112 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_112, 0, x_89); -lean_ctor_set(x_112, 1, x_111); -x_113 = 0; -x_114 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_112, x_113, x_1, x_2, x_3, x_4, x_66); +lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; +lean_dec(x_89); +x_109 = l_Lean_Meta_reportDiag___closed__15; +x_110 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_110, 0, x_87); +lean_ctor_set(x_110, 1, x_109); +x_111 = 0; +x_112 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_110, x_111, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_114; +return x_112; } } case 5: { -uint8_t x_115; -lean_dec(x_90); -lean_free_object(x_63); -x_115 = !lean_is_exclusive(x_91); -if (x_115 == 0) -{ -lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; -x_116 = lean_ctor_get(x_91, 1); -lean_dec(x_116); -x_117 = lean_ctor_get(x_91, 0); -lean_dec(x_117); -x_118 = l_Lean_Meta_reportDiag___closed__16; -lean_ctor_set_tag(x_91, 7); -lean_ctor_set(x_91, 1, x_118); -lean_ctor_set(x_91, 0, x_89); -x_119 = 0; -x_120 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_91, x_119, x_1, x_2, x_3, x_4, x_66); +uint8_t x_113; +lean_dec(x_88); +lean_free_object(x_61); +x_113 = !lean_is_exclusive(x_89); +if (x_113 == 0) +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; uint8_t x_117; lean_object* x_118; +x_114 = lean_ctor_get(x_89, 1); +lean_dec(x_114); +x_115 = lean_ctor_get(x_89, 0); +lean_dec(x_115); +x_116 = l_Lean_Meta_reportDiag___closed__15; +lean_ctor_set_tag(x_89, 7); +lean_ctor_set(x_89, 1, x_116); +lean_ctor_set(x_89, 0, x_87); +x_117 = 0; +x_118 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_89, x_117, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_120; +return x_118; } else { -lean_object* x_121; lean_object* x_122; uint8_t x_123; lean_object* x_124; -lean_dec(x_91); -x_121 = l_Lean_Meta_reportDiag___closed__16; -x_122 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_122, 0, x_89); -lean_ctor_set(x_122, 1, x_121); -x_123 = 0; -x_124 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_122, x_123, x_1, x_2, x_3, x_4, x_66); +lean_object* x_119; lean_object* x_120; uint8_t x_121; lean_object* x_122; +lean_dec(x_89); +x_119 = l_Lean_Meta_reportDiag___closed__15; +x_120 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_120, 0, x_87); +lean_ctor_set(x_120, 1, x_119); +x_121 = 0; +x_122 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_120, x_121, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_124; +return x_122; } } case 7: { -uint8_t x_125; -lean_dec(x_90); -lean_free_object(x_63); -x_125 = !lean_is_exclusive(x_91); -if (x_125 == 0) -{ -lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; -x_126 = lean_ctor_get(x_91, 1); -lean_dec(x_126); -x_127 = lean_ctor_get(x_91, 0); -lean_dec(x_127); -x_128 = l_Lean_Meta_reportDiag___closed__16; -lean_ctor_set(x_91, 1, x_128); -lean_ctor_set(x_91, 0, x_89); -x_129 = 0; -x_130 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_91, x_129, x_1, x_2, x_3, x_4, x_66); +uint8_t x_123; +lean_dec(x_88); +lean_free_object(x_61); +x_123 = !lean_is_exclusive(x_89); +if (x_123 == 0) +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; lean_object* x_128; +x_124 = lean_ctor_get(x_89, 1); +lean_dec(x_124); +x_125 = lean_ctor_get(x_89, 0); +lean_dec(x_125); +x_126 = l_Lean_Meta_reportDiag___closed__15; +lean_ctor_set(x_89, 1, x_126); +lean_ctor_set(x_89, 0, x_87); +x_127 = 0; +x_128 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_89, x_127, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_130; +return x_128; } else { -lean_object* x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; -lean_dec(x_91); -x_131 = l_Lean_Meta_reportDiag___closed__16; -x_132 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_132, 0, x_89); -lean_ctor_set(x_132, 1, x_131); -x_133 = 0; -x_134 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_132, x_133, x_1, x_2, x_3, x_4, x_66); +lean_object* x_129; lean_object* x_130; uint8_t x_131; lean_object* x_132; +lean_dec(x_89); +x_129 = l_Lean_Meta_reportDiag___closed__15; +x_130 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_130, 0, x_87); +lean_ctor_set(x_130, 1, x_129); +x_131 = 0; +x_132 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_130, x_131, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_134; +return x_132; } } default: { -uint8_t x_135; -lean_dec(x_91); -lean_free_object(x_63); -x_135 = !lean_is_exclusive(x_90); -if (x_135 == 0) -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_140; -x_136 = lean_ctor_get(x_90, 1); -lean_dec(x_136); -x_137 = lean_ctor_get(x_90, 0); -lean_dec(x_137); -x_138 = l_Lean_Meta_reportDiag___closed__16; -lean_ctor_set_tag(x_90, 7); -lean_ctor_set(x_90, 1, x_138); -lean_ctor_set(x_90, 0, x_89); -x_139 = 0; -x_140 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_90, x_139, x_1, x_2, x_3, x_4, x_66); +uint8_t x_133; +lean_dec(x_89); +lean_free_object(x_61); +x_133 = !lean_is_exclusive(x_88); +if (x_133 == 0) +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; lean_object* x_138; +x_134 = lean_ctor_get(x_88, 1); +lean_dec(x_134); +x_135 = lean_ctor_get(x_88, 0); +lean_dec(x_135); +x_136 = l_Lean_Meta_reportDiag___closed__15; +lean_ctor_set_tag(x_88, 7); +lean_ctor_set(x_88, 1, x_136); +lean_ctor_set(x_88, 0, x_87); +x_137 = 0; +x_138 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_88, x_137, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_140; +return x_138; } else { -lean_object* x_141; lean_object* x_142; uint8_t x_143; lean_object* x_144; -lean_dec(x_90); -x_141 = l_Lean_Meta_reportDiag___closed__16; -x_142 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_142, 0, x_89); -lean_ctor_set(x_142, 1, x_141); -x_143 = 0; -x_144 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_142, x_143, x_1, x_2, x_3, x_4, x_66); +lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_object* x_142; +lean_dec(x_88); +x_139 = l_Lean_Meta_reportDiag___closed__15; +x_140 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_140, 0, x_87); +lean_ctor_set(x_140, 1, x_139); +x_141 = 0; +x_142 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_140, x_141, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_144; +return x_142; } } } } else { -lean_object* x_145; uint8_t x_146; lean_object* x_147; -lean_free_object(x_63); -x_145 = l_Lean_Meta_reportDiag___closed__16; +lean_object* x_143; uint8_t x_144; lean_object* x_145; +lean_free_object(x_61); +x_143 = l_Lean_Meta_reportDiag___closed__15; lean_ctor_set_tag(x_53, 7); -lean_ctor_set(x_53, 1, x_145); -lean_ctor_set(x_53, 0, x_89); -x_146 = 0; -x_147 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_53, x_146, x_1, x_2, x_3, x_4, x_66); +lean_ctor_set(x_53, 1, x_143); +lean_ctor_set(x_53, 0, x_87); +x_144 = 0; +x_145 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_53, x_144, x_1, x_2, x_3, x_4, x_64); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_147; +return x_145; } } else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_148 = lean_ctor_get(x_63, 0); -x_149 = lean_ctor_get(x_63, 1); -lean_inc(x_149); -lean_inc(x_148); -lean_dec(x_63); -x_150 = l_Lean_MessageData_nil; -x_151 = l_Lean_Meta_mkDiagSummaryForUnfolded___closed__2; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_146 = lean_ctor_get(x_61, 0); +x_147 = lean_ctor_get(x_61, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_61); +x_148 = l_Lean_MessageData_nil; +x_149 = l_Lean_Meta_mkDiagSummaryForUnfolded___closed__2; +x_150 = l_Lean_Meta_reportDiag___closed__5; +x_151 = l_Lean_Meta_appendSection(x_148, x_149, x_150, x_23, x_25); x_152 = l_Lean_Meta_reportDiag___closed__6; -x_153 = l_Lean_Meta_appendSection(x_150, x_151, x_152, x_23, x_25); +x_153 = l_Lean_Meta_appendSection(x_151, x_149, x_152, x_27, x_25); x_154 = l_Lean_Meta_reportDiag___closed__7; -x_155 = l_Lean_Meta_appendSection(x_153, x_151, x_154, x_27, x_25); -x_156 = l_Lean_Meta_reportDiag___closed__8; -x_157 = l_Lean_Meta_appendSection(x_155, x_151, x_156, x_30, x_25); -x_158 = l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__2; -x_159 = l_Lean_Meta_reportDiag___closed__9; -x_160 = l_Lean_Meta_appendSection(x_157, x_158, x_159, x_43, x_25); -x_161 = lean_ctor_get(x_3, 2); -lean_inc(x_161); -x_162 = l_Lean_Meta_reportDiag___closed__10; -x_163 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_161, x_162); -lean_dec(x_161); -x_164 = l___private_Init_Data_Repr_0__Nat_reprFast(x_163); +x_155 = l_Lean_Meta_appendSection(x_153, x_149, x_154, x_30, x_25); +x_156 = l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__2; +x_157 = l_Lean_Meta_reportDiag___closed__8; +x_158 = l_Lean_Meta_appendSection(x_155, x_156, x_157, x_43, x_25); +x_159 = lean_ctor_get(x_3, 2); +lean_inc(x_159); +x_160 = l_Lean_Meta_reportDiag___closed__9; +x_161 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_159, x_160); +lean_dec(x_159); +x_162 = l___private_Init_Data_Repr_0__Nat_reprFast(x_161); +x_163 = l_Lean_Meta_reportDiag___closed__10; +x_164 = lean_string_append(x_163, x_162); +lean_dec(x_162); x_165 = l_Lean_Meta_reportDiag___closed__11; -x_166 = lean_string_append(x_165, x_164); -lean_dec(x_164); -x_167 = l_Lean_Meta_reportDiag___closed__12; -x_168 = lean_string_append(x_166, x_167); -x_169 = l_Lean_Meta_appendSection(x_160, x_158, x_168, x_51, x_21); -x_170 = l_Lean_Meta_reportDiag___closed__13; -x_171 = l_Lean_Meta_appendSection(x_169, x_37, x_170, x_40, x_25); -x_172 = l_Lean_Meta_appendSection(x_171, x_62, x_152, x_148, x_25); -if (lean_obj_tag(x_172) == 0) -{ -lean_object* x_173; lean_object* x_174; +x_166 = lean_string_append(x_164, x_165); +x_167 = l_Lean_Meta_appendSection(x_158, x_156, x_166, x_51, x_21); +x_168 = l_Lean_Meta_reportDiag___closed__12; +x_169 = l_Lean_Meta_appendSection(x_167, x_37, x_168, x_40, x_25); +x_170 = l_Lean_Meta_appendSection(x_169, x_60, x_150, x_146, x_25); +if (lean_obj_tag(x_170) == 0) +{ +lean_object* x_171; lean_object* x_172; lean_free_object(x_53); -x_173 = lean_ctor_get(x_172, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_173, 0); -lean_inc(x_174); -switch (lean_obj_tag(x_174)) { +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +switch (lean_obj_tag(x_172)) { case 0: { -lean_object* x_175; lean_object* x_176; -x_175 = lean_ctor_get(x_173, 1); -lean_inc(x_175); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_176 = x_173; +lean_object* x_173; lean_object* x_174; +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_174 = x_171; } else { - lean_dec_ref(x_173); - x_176 = lean_box(0); + lean_dec_ref(x_171); + x_174 = lean_box(0); } -if (lean_obj_tag(x_175) == 0) +if (lean_obj_tag(x_173) == 0) { -lean_object* x_177; lean_object* x_178; -lean_dec(x_176); -lean_dec(x_172); +lean_object* x_175; lean_object* x_176; +lean_dec(x_174); +lean_dec(x_170); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_177 = lean_box(0); -x_178 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_178, 0, x_177); -lean_ctor_set(x_178, 1, x_149); -return x_178; +x_175 = lean_box(0); +x_176 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_176, 0, x_175); +lean_ctor_set(x_176, 1, x_147); +return x_176; } else { -lean_object* x_179; lean_object* x_180; uint8_t x_181; lean_object* x_182; -lean_dec(x_175); -x_179 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_176)) { - x_180 = lean_alloc_ctor(7, 2, 0); +lean_object* x_177; lean_object* x_178; uint8_t x_179; lean_object* x_180; +lean_dec(x_173); +x_177 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_174)) { + x_178 = lean_alloc_ctor(7, 2, 0); } else { - x_180 = x_176; - lean_ctor_set_tag(x_180, 7); + x_178 = x_174; + lean_ctor_set_tag(x_178, 7); } -lean_ctor_set(x_180, 0, x_172); -lean_ctor_set(x_180, 1, x_179); -x_181 = 0; -x_182 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_180, x_181, x_1, x_2, x_3, x_4, x_149); +lean_ctor_set(x_178, 0, x_170); +lean_ctor_set(x_178, 1, x_177); +x_179 = 0; +x_180 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_178, x_179, x_1, x_2, x_3, x_4, x_147); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_182; +return x_180; } } case 4: { -lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; lean_object* x_187; -lean_dec(x_173); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_183 = x_174; +lean_object* x_181; lean_object* x_182; lean_object* x_183; uint8_t x_184; lean_object* x_185; +lean_dec(x_171); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + x_181 = x_172; } else { - lean_dec_ref(x_174); - x_183 = lean_box(0); + lean_dec_ref(x_172); + x_181 = lean_box(0); } -x_184 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_183)) { - x_185 = lean_alloc_ctor(7, 2, 0); +x_182 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_181)) { + x_183 = lean_alloc_ctor(7, 2, 0); } else { - x_185 = x_183; - lean_ctor_set_tag(x_185, 7); + x_183 = x_181; + lean_ctor_set_tag(x_183, 7); } -lean_ctor_set(x_185, 0, x_172); -lean_ctor_set(x_185, 1, x_184); -x_186 = 0; -x_187 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_185, x_186, x_1, x_2, x_3, x_4, x_149); +lean_ctor_set(x_183, 0, x_170); +lean_ctor_set(x_183, 1, x_182); +x_184 = 0; +x_185 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_183, x_184, x_1, x_2, x_3, x_4, x_147); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_187; +return x_185; } case 5: { -lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; lean_object* x_192; -lean_dec(x_173); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_188 = x_174; +lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; +lean_dec(x_171); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + x_186 = x_172; } else { - lean_dec_ref(x_174); - x_188 = lean_box(0); + lean_dec_ref(x_172); + x_186 = lean_box(0); } -x_189 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_188)) { - x_190 = lean_alloc_ctor(7, 2, 0); +x_187 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_186)) { + x_188 = lean_alloc_ctor(7, 2, 0); } else { - x_190 = x_188; - lean_ctor_set_tag(x_190, 7); + x_188 = x_186; + lean_ctor_set_tag(x_188, 7); } -lean_ctor_set(x_190, 0, x_172); -lean_ctor_set(x_190, 1, x_189); -x_191 = 0; -x_192 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_190, x_191, x_1, x_2, x_3, x_4, x_149); +lean_ctor_set(x_188, 0, x_170); +lean_ctor_set(x_188, 1, x_187); +x_189 = 0; +x_190 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_188, x_189, x_1, x_2, x_3, x_4, x_147); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_192; +return x_190; } case 7: { -lean_object* x_193; lean_object* x_194; lean_object* x_195; uint8_t x_196; lean_object* x_197; -lean_dec(x_173); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_193 = x_174; +lean_object* x_191; lean_object* x_192; lean_object* x_193; uint8_t x_194; lean_object* x_195; +lean_dec(x_171); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + x_191 = x_172; } else { - lean_dec_ref(x_174); - x_193 = lean_box(0); + lean_dec_ref(x_172); + x_191 = lean_box(0); } -x_194 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_193)) { - x_195 = lean_alloc_ctor(7, 2, 0); +x_192 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_191)) { + x_193 = lean_alloc_ctor(7, 2, 0); } else { - x_195 = x_193; + x_193 = x_191; } -lean_ctor_set(x_195, 0, x_172); -lean_ctor_set(x_195, 1, x_194); -x_196 = 0; -x_197 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_195, x_196, x_1, x_2, x_3, x_4, x_149); +lean_ctor_set(x_193, 0, x_170); +lean_ctor_set(x_193, 1, x_192); +x_194 = 0; +x_195 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_193, x_194, x_1, x_2, x_3, x_4, x_147); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_197; +return x_195; } default: { -lean_object* x_198; lean_object* x_199; lean_object* x_200; uint8_t x_201; lean_object* x_202; -lean_dec(x_174); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_198 = x_173; +lean_object* x_196; lean_object* x_197; lean_object* x_198; uint8_t x_199; lean_object* x_200; +lean_dec(x_172); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_196 = x_171; } else { - lean_dec_ref(x_173); - x_198 = lean_box(0); + lean_dec_ref(x_171); + x_196 = lean_box(0); } -x_199 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_198)) { - x_200 = lean_alloc_ctor(7, 2, 0); +x_197 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_196)) { + x_198 = lean_alloc_ctor(7, 2, 0); } else { - x_200 = x_198; - lean_ctor_set_tag(x_200, 7); + x_198 = x_196; + lean_ctor_set_tag(x_198, 7); } -lean_ctor_set(x_200, 0, x_172); -lean_ctor_set(x_200, 1, x_199); -x_201 = 0; -x_202 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_200, x_201, x_1, x_2, x_3, x_4, x_149); +lean_ctor_set(x_198, 0, x_170); +lean_ctor_set(x_198, 1, x_197); +x_199 = 0; +x_200 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_198, x_199, x_1, x_2, x_3, x_4, x_147); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_202; +return x_200; } } } else { -lean_object* x_203; uint8_t x_204; lean_object* x_205; -x_203 = l_Lean_Meta_reportDiag___closed__16; +lean_object* x_201; uint8_t x_202; lean_object* x_203; +x_201 = l_Lean_Meta_reportDiag___closed__15; lean_ctor_set_tag(x_53, 7); -lean_ctor_set(x_53, 1, x_203); -lean_ctor_set(x_53, 0, x_172); -x_204 = 0; -x_205 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_53, x_204, x_1, x_2, x_3, x_4, x_149); +lean_ctor_set(x_53, 1, x_201); +lean_ctor_set(x_53, 0, x_170); +x_202 = 0; +x_203 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_53, x_202, x_1, x_2, x_3, x_4, x_147); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_205; +return x_203; } } } else { -uint8_t x_206; +uint8_t x_204; lean_free_object(x_53); lean_dec(x_51); lean_dec(x_43); @@ -6001,285 +5988,284 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_206 = !lean_is_exclusive(x_63); -if (x_206 == 0) +x_204 = !lean_is_exclusive(x_61); +if (x_204 == 0) { -return x_63; +return x_61; } else { -lean_object* x_207; lean_object* x_208; lean_object* x_209; -x_207 = lean_ctor_get(x_63, 0); -x_208 = lean_ctor_get(x_63, 1); -lean_inc(x_208); -lean_inc(x_207); -lean_dec(x_63); -x_209 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_209, 0, x_207); -lean_ctor_set(x_209, 1, x_208); -return x_209; +lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_205 = lean_ctor_get(x_61, 0); +x_206 = lean_ctor_get(x_61, 1); +lean_inc(x_206); +lean_inc(x_205); +lean_dec(x_61); +x_207 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_207, 0, x_205); +lean_ctor_set(x_207, 1, x_206); +return x_207; } } } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; -x_210 = lean_ctor_get(x_53, 0); -x_211 = lean_ctor_get(x_53, 1); -lean_inc(x_211); -lean_inc(x_210); +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_208 = lean_ctor_get(x_53, 0); +x_209 = lean_ctor_get(x_53, 1); +lean_inc(x_209); +lean_inc(x_208); lean_dec(x_53); -x_212 = lean_ctor_get(x_210, 0); -lean_inc(x_212); +x_210 = lean_ctor_get(x_208, 0); +lean_inc(x_210); +lean_dec(x_208); +x_211 = lean_ctor_get(x_210, 1); +lean_inc(x_211); lean_dec(x_210); -x_213 = l_Lean_Kernel_instInhabitedDiagnostics; -x_214 = l_Lean_Meta_reportDiag___closed__5; -x_215 = l_Lean_EnvExtension_getState___rarg(x_213, x_214, x_212); -lean_dec(x_212); -x_216 = lean_ctor_get(x_215, 0); +x_212 = lean_ctor_get(x_211, 0); +lean_inc(x_212); +lean_dec(x_211); +x_213 = l_Lean_Meta_reportDiag___closed__4; +x_214 = l_Lean_Meta_mkDiagSummary(x_213, x_212, x_38, x_1, x_2, x_3, x_4, x_209); +if (lean_obj_tag(x_214) == 0) +{ +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_214, 1); lean_inc(x_216); -lean_dec(x_215); -x_217 = l_Lean_Meta_reportDiag___closed__4; -x_218 = l_Lean_Meta_mkDiagSummary(x_217, x_216, x_38, x_1, x_2, x_3, x_4, x_211); -if (lean_obj_tag(x_218) == 0) -{ -lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; -x_219 = lean_ctor_get(x_218, 0); -lean_inc(x_219); -x_220 = lean_ctor_get(x_218, 1); -lean_inc(x_220); -if (lean_is_exclusive(x_218)) { - lean_ctor_release(x_218, 0); - lean_ctor_release(x_218, 1); - x_221 = x_218; +if (lean_is_exclusive(x_214)) { + lean_ctor_release(x_214, 0); + lean_ctor_release(x_214, 1); + x_217 = x_214; } else { - lean_dec_ref(x_218); - x_221 = lean_box(0); -} -x_222 = l_Lean_MessageData_nil; -x_223 = l_Lean_Meta_mkDiagSummaryForUnfolded___closed__2; -x_224 = l_Lean_Meta_reportDiag___closed__6; -x_225 = l_Lean_Meta_appendSection(x_222, x_223, x_224, x_23, x_25); -x_226 = l_Lean_Meta_reportDiag___closed__7; -x_227 = l_Lean_Meta_appendSection(x_225, x_223, x_226, x_27, x_25); -x_228 = l_Lean_Meta_reportDiag___closed__8; -x_229 = l_Lean_Meta_appendSection(x_227, x_223, x_228, x_30, x_25); -x_230 = l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__2; -x_231 = l_Lean_Meta_reportDiag___closed__9; -x_232 = l_Lean_Meta_appendSection(x_229, x_230, x_231, x_43, x_25); -x_233 = lean_ctor_get(x_3, 2); -lean_inc(x_233); -x_234 = l_Lean_Meta_reportDiag___closed__10; -x_235 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_233, x_234); -lean_dec(x_233); -x_236 = l___private_Init_Data_Repr_0__Nat_reprFast(x_235); -x_237 = l_Lean_Meta_reportDiag___closed__11; -x_238 = lean_string_append(x_237, x_236); -lean_dec(x_236); -x_239 = l_Lean_Meta_reportDiag___closed__12; -x_240 = lean_string_append(x_238, x_239); -x_241 = l_Lean_Meta_appendSection(x_232, x_230, x_240, x_51, x_21); -x_242 = l_Lean_Meta_reportDiag___closed__13; -x_243 = l_Lean_Meta_appendSection(x_241, x_37, x_242, x_40, x_25); -x_244 = l_Lean_Meta_appendSection(x_243, x_217, x_224, x_219, x_25); -if (lean_obj_tag(x_244) == 0) -{ -lean_object* x_245; lean_object* x_246; -x_245 = lean_ctor_get(x_244, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_245, 0); -lean_inc(x_246); -switch (lean_obj_tag(x_246)) { + lean_dec_ref(x_214); + x_217 = lean_box(0); +} +x_218 = l_Lean_MessageData_nil; +x_219 = l_Lean_Meta_mkDiagSummaryForUnfolded___closed__2; +x_220 = l_Lean_Meta_reportDiag___closed__5; +x_221 = l_Lean_Meta_appendSection(x_218, x_219, x_220, x_23, x_25); +x_222 = l_Lean_Meta_reportDiag___closed__6; +x_223 = l_Lean_Meta_appendSection(x_221, x_219, x_222, x_27, x_25); +x_224 = l_Lean_Meta_reportDiag___closed__7; +x_225 = l_Lean_Meta_appendSection(x_223, x_219, x_224, x_30, x_25); +x_226 = l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__2; +x_227 = l_Lean_Meta_reportDiag___closed__8; +x_228 = l_Lean_Meta_appendSection(x_225, x_226, x_227, x_43, x_25); +x_229 = lean_ctor_get(x_3, 2); +lean_inc(x_229); +x_230 = l_Lean_Meta_reportDiag___closed__9; +x_231 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_229, x_230); +lean_dec(x_229); +x_232 = l___private_Init_Data_Repr_0__Nat_reprFast(x_231); +x_233 = l_Lean_Meta_reportDiag___closed__10; +x_234 = lean_string_append(x_233, x_232); +lean_dec(x_232); +x_235 = l_Lean_Meta_reportDiag___closed__11; +x_236 = lean_string_append(x_234, x_235); +x_237 = l_Lean_Meta_appendSection(x_228, x_226, x_236, x_51, x_21); +x_238 = l_Lean_Meta_reportDiag___closed__12; +x_239 = l_Lean_Meta_appendSection(x_237, x_37, x_238, x_40, x_25); +x_240 = l_Lean_Meta_appendSection(x_239, x_213, x_220, x_215, x_25); +if (lean_obj_tag(x_240) == 0) +{ +lean_object* x_241; lean_object* x_242; +x_241 = lean_ctor_get(x_240, 0); +lean_inc(x_241); +x_242 = lean_ctor_get(x_241, 0); +lean_inc(x_242); +switch (lean_obj_tag(x_242)) { case 0: { -lean_object* x_247; lean_object* x_248; -x_247 = lean_ctor_get(x_245, 1); -lean_inc(x_247); -if (lean_is_exclusive(x_245)) { - lean_ctor_release(x_245, 0); - lean_ctor_release(x_245, 1); - x_248 = x_245; +lean_object* x_243; lean_object* x_244; +x_243 = lean_ctor_get(x_241, 1); +lean_inc(x_243); +if (lean_is_exclusive(x_241)) { + lean_ctor_release(x_241, 0); + lean_ctor_release(x_241, 1); + x_244 = x_241; } else { - lean_dec_ref(x_245); - x_248 = lean_box(0); + lean_dec_ref(x_241); + x_244 = lean_box(0); } -if (lean_obj_tag(x_247) == 0) +if (lean_obj_tag(x_243) == 0) { -lean_object* x_249; lean_object* x_250; -lean_dec(x_248); +lean_object* x_245; lean_object* x_246; lean_dec(x_244); +lean_dec(x_240); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_249 = lean_box(0); -if (lean_is_scalar(x_221)) { - x_250 = lean_alloc_ctor(0, 2, 0); +x_245 = lean_box(0); +if (lean_is_scalar(x_217)) { + x_246 = lean_alloc_ctor(0, 2, 0); } else { - x_250 = x_221; + x_246 = x_217; } -lean_ctor_set(x_250, 0, x_249); -lean_ctor_set(x_250, 1, x_220); -return x_250; +lean_ctor_set(x_246, 0, x_245); +lean_ctor_set(x_246, 1, x_216); +return x_246; } else { -lean_object* x_251; lean_object* x_252; uint8_t x_253; lean_object* x_254; -lean_dec(x_247); -lean_dec(x_221); -x_251 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_248)) { - x_252 = lean_alloc_ctor(7, 2, 0); +lean_object* x_247; lean_object* x_248; uint8_t x_249; lean_object* x_250; +lean_dec(x_243); +lean_dec(x_217); +x_247 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_244)) { + x_248 = lean_alloc_ctor(7, 2, 0); } else { - x_252 = x_248; - lean_ctor_set_tag(x_252, 7); + x_248 = x_244; + lean_ctor_set_tag(x_248, 7); } -lean_ctor_set(x_252, 0, x_244); -lean_ctor_set(x_252, 1, x_251); -x_253 = 0; -x_254 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_252, x_253, x_1, x_2, x_3, x_4, x_220); +lean_ctor_set(x_248, 0, x_240); +lean_ctor_set(x_248, 1, x_247); +x_249 = 0; +x_250 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_248, x_249, x_1, x_2, x_3, x_4, x_216); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_254; +return x_250; } } case 4: { -lean_object* x_255; lean_object* x_256; lean_object* x_257; uint8_t x_258; lean_object* x_259; -lean_dec(x_245); -lean_dec(x_221); -if (lean_is_exclusive(x_246)) { - lean_ctor_release(x_246, 0); - lean_ctor_release(x_246, 1); - x_255 = x_246; +lean_object* x_251; lean_object* x_252; lean_object* x_253; uint8_t x_254; lean_object* x_255; +lean_dec(x_241); +lean_dec(x_217); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_251 = x_242; } else { - lean_dec_ref(x_246); - x_255 = lean_box(0); + lean_dec_ref(x_242); + x_251 = lean_box(0); } -x_256 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_255)) { - x_257 = lean_alloc_ctor(7, 2, 0); +x_252 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_251)) { + x_253 = lean_alloc_ctor(7, 2, 0); } else { - x_257 = x_255; - lean_ctor_set_tag(x_257, 7); + x_253 = x_251; + lean_ctor_set_tag(x_253, 7); } -lean_ctor_set(x_257, 0, x_244); -lean_ctor_set(x_257, 1, x_256); -x_258 = 0; -x_259 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_257, x_258, x_1, x_2, x_3, x_4, x_220); +lean_ctor_set(x_253, 0, x_240); +lean_ctor_set(x_253, 1, x_252); +x_254 = 0; +x_255 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_253, x_254, x_1, x_2, x_3, x_4, x_216); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_259; +return x_255; } case 5: { -lean_object* x_260; lean_object* x_261; lean_object* x_262; uint8_t x_263; lean_object* x_264; -lean_dec(x_245); -lean_dec(x_221); -if (lean_is_exclusive(x_246)) { - lean_ctor_release(x_246, 0); - lean_ctor_release(x_246, 1); - x_260 = x_246; +lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; lean_object* x_260; +lean_dec(x_241); +lean_dec(x_217); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_256 = x_242; } else { - lean_dec_ref(x_246); - x_260 = lean_box(0); + lean_dec_ref(x_242); + x_256 = lean_box(0); } -x_261 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_260)) { - x_262 = lean_alloc_ctor(7, 2, 0); +x_257 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_256)) { + x_258 = lean_alloc_ctor(7, 2, 0); } else { - x_262 = x_260; - lean_ctor_set_tag(x_262, 7); + x_258 = x_256; + lean_ctor_set_tag(x_258, 7); } -lean_ctor_set(x_262, 0, x_244); -lean_ctor_set(x_262, 1, x_261); -x_263 = 0; -x_264 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_262, x_263, x_1, x_2, x_3, x_4, x_220); +lean_ctor_set(x_258, 0, x_240); +lean_ctor_set(x_258, 1, x_257); +x_259 = 0; +x_260 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_258, x_259, x_1, x_2, x_3, x_4, x_216); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_264; +return x_260; } case 7: { -lean_object* x_265; lean_object* x_266; lean_object* x_267; uint8_t x_268; lean_object* x_269; -lean_dec(x_245); -lean_dec(x_221); -if (lean_is_exclusive(x_246)) { - lean_ctor_release(x_246, 0); - lean_ctor_release(x_246, 1); - x_265 = x_246; +lean_object* x_261; lean_object* x_262; lean_object* x_263; uint8_t x_264; lean_object* x_265; +lean_dec(x_241); +lean_dec(x_217); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_261 = x_242; } else { - lean_dec_ref(x_246); - x_265 = lean_box(0); + lean_dec_ref(x_242); + x_261 = lean_box(0); } -x_266 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_265)) { - x_267 = lean_alloc_ctor(7, 2, 0); +x_262 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_261)) { + x_263 = lean_alloc_ctor(7, 2, 0); } else { - x_267 = x_265; + x_263 = x_261; } -lean_ctor_set(x_267, 0, x_244); -lean_ctor_set(x_267, 1, x_266); -x_268 = 0; -x_269 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_267, x_268, x_1, x_2, x_3, x_4, x_220); +lean_ctor_set(x_263, 0, x_240); +lean_ctor_set(x_263, 1, x_262); +x_264 = 0; +x_265 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_263, x_264, x_1, x_2, x_3, x_4, x_216); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_269; +return x_265; } default: { -lean_object* x_270; lean_object* x_271; lean_object* x_272; uint8_t x_273; lean_object* x_274; -lean_dec(x_246); -lean_dec(x_221); -if (lean_is_exclusive(x_245)) { - lean_ctor_release(x_245, 0); - lean_ctor_release(x_245, 1); - x_270 = x_245; +lean_object* x_266; lean_object* x_267; lean_object* x_268; uint8_t x_269; lean_object* x_270; +lean_dec(x_242); +lean_dec(x_217); +if (lean_is_exclusive(x_241)) { + lean_ctor_release(x_241, 0); + lean_ctor_release(x_241, 1); + x_266 = x_241; } else { - lean_dec_ref(x_245); - x_270 = lean_box(0); + lean_dec_ref(x_241); + x_266 = lean_box(0); } -x_271 = l_Lean_Meta_reportDiag___closed__16; -if (lean_is_scalar(x_270)) { - x_272 = lean_alloc_ctor(7, 2, 0); +x_267 = l_Lean_Meta_reportDiag___closed__15; +if (lean_is_scalar(x_266)) { + x_268 = lean_alloc_ctor(7, 2, 0); } else { - x_272 = x_270; - lean_ctor_set_tag(x_272, 7); + x_268 = x_266; + lean_ctor_set_tag(x_268, 7); } -lean_ctor_set(x_272, 0, x_244); -lean_ctor_set(x_272, 1, x_271); -x_273 = 0; -x_274 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_272, x_273, x_1, x_2, x_3, x_4, x_220); +lean_ctor_set(x_268, 0, x_240); +lean_ctor_set(x_268, 1, x_267); +x_269 = 0; +x_270 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_268, x_269, x_1, x_2, x_3, x_4, x_216); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_274; +return x_270; } } } else { -lean_object* x_275; lean_object* x_276; uint8_t x_277; lean_object* x_278; -lean_dec(x_221); -x_275 = l_Lean_Meta_reportDiag___closed__16; -x_276 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_276, 0, x_244); -lean_ctor_set(x_276, 1, x_275); -x_277 = 0; -x_278 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_276, x_277, x_1, x_2, x_3, x_4, x_220); +lean_object* x_271; lean_object* x_272; uint8_t x_273; lean_object* x_274; +lean_dec(x_217); +x_271 = l_Lean_Meta_reportDiag___closed__15; +x_272 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_272, 0, x_240); +lean_ctor_set(x_272, 1, x_271); +x_273 = 0; +x_274 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_272, x_273, x_1, x_2, x_3, x_4, x_216); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_278; +return x_274; } } else { -lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_dec(x_51); lean_dec(x_43); lean_dec(x_40); @@ -6290,25 +6276,57 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_279 = lean_ctor_get(x_218, 0); -lean_inc(x_279); -x_280 = lean_ctor_get(x_218, 1); -lean_inc(x_280); -if (lean_is_exclusive(x_218)) { - lean_ctor_release(x_218, 0); - lean_ctor_release(x_218, 1); - x_281 = x_218; +x_275 = lean_ctor_get(x_214, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_214, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_214)) { + lean_ctor_release(x_214, 0); + lean_ctor_release(x_214, 1); + x_277 = x_214; } else { - lean_dec_ref(x_218); - x_281 = lean_box(0); + lean_dec_ref(x_214); + x_277 = lean_box(0); } -if (lean_is_scalar(x_281)) { - x_282 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_277)) { + x_278 = lean_alloc_ctor(1, 2, 0); } else { - x_282 = x_281; + x_278 = x_277; +} +lean_ctor_set(x_278, 0, x_275); +lean_ctor_set(x_278, 1, x_276); +return x_278; +} +} } -lean_ctor_set(x_282, 0, x_279); -lean_ctor_set(x_282, 1, x_280); +else +{ +uint8_t x_279; +lean_dec(x_43); +lean_dec(x_40); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_279 = !lean_is_exclusive(x_50); +if (x_279 == 0) +{ +return x_50; +} +else +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; +x_280 = lean_ctor_get(x_50, 0); +x_281 = lean_ctor_get(x_50, 1); +lean_inc(x_281); +lean_inc(x_280); +lean_dec(x_50); +x_282 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_282, 0, x_280); +lean_ctor_set(x_282, 1, x_281); return x_282; } } @@ -6316,7 +6334,6 @@ return x_282; else { uint8_t x_283; -lean_dec(x_43); lean_dec(x_40); lean_dec(x_30); lean_dec(x_27); @@ -6325,19 +6342,19 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_283 = !lean_is_exclusive(x_50); +x_283 = !lean_is_exclusive(x_42); if (x_283 == 0) { -return x_50; +return x_42; } else { lean_object* x_284; lean_object* x_285; lean_object* x_286; -x_284 = lean_ctor_get(x_50, 0); -x_285 = lean_ctor_get(x_50, 1); +x_284 = lean_ctor_get(x_42, 0); +x_285 = lean_ctor_get(x_42, 1); lean_inc(x_285); lean_inc(x_284); -lean_dec(x_50); +lean_dec(x_42); x_286 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_286, 0, x_284); lean_ctor_set(x_286, 1, x_285); @@ -6348,7 +6365,6 @@ return x_286; else { uint8_t x_287; -lean_dec(x_40); lean_dec(x_30); lean_dec(x_27); lean_dec(x_23); @@ -6356,19 +6372,19 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_287 = !lean_is_exclusive(x_42); +x_287 = !lean_is_exclusive(x_39); if (x_287 == 0) { -return x_42; +return x_39; } else { lean_object* x_288; lean_object* x_289; lean_object* x_290; -x_288 = lean_ctor_get(x_42, 0); -x_289 = lean_ctor_get(x_42, 1); +x_288 = lean_ctor_get(x_39, 0); +x_289 = lean_ctor_get(x_39, 1); lean_inc(x_289); lean_inc(x_288); -lean_dec(x_42); +lean_dec(x_39); x_290 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_290, 0, x_288); lean_ctor_set(x_290, 1, x_289); @@ -6379,26 +6395,25 @@ return x_290; else { uint8_t x_291; -lean_dec(x_30); lean_dec(x_27); lean_dec(x_23); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_291 = !lean_is_exclusive(x_39); +x_291 = !lean_is_exclusive(x_29); if (x_291 == 0) { -return x_39; +return x_29; } else { lean_object* x_292; lean_object* x_293; lean_object* x_294; -x_292 = lean_ctor_get(x_39, 0); -x_293 = lean_ctor_get(x_39, 1); +x_292 = lean_ctor_get(x_29, 0); +x_293 = lean_ctor_get(x_29, 1); lean_inc(x_293); lean_inc(x_292); -lean_dec(x_39); +lean_dec(x_29); x_294 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_294, 0, x_292); lean_ctor_set(x_294, 1, x_293); @@ -6409,25 +6424,25 @@ return x_294; else { uint8_t x_295; -lean_dec(x_27); lean_dec(x_23); +lean_dec(x_20); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_295 = !lean_is_exclusive(x_29); +x_295 = !lean_is_exclusive(x_26); if (x_295 == 0) { -return x_29; +return x_26; } else { lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_296 = lean_ctor_get(x_29, 0); -x_297 = lean_ctor_get(x_29, 1); +x_296 = lean_ctor_get(x_26, 0); +x_297 = lean_ctor_get(x_26, 1); lean_inc(x_297); lean_inc(x_296); -lean_dec(x_29); +lean_dec(x_26); x_298 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_298, 0, x_296); lean_ctor_set(x_298, 1, x_297); @@ -6438,25 +6453,24 @@ return x_298; else { uint8_t x_299; -lean_dec(x_23); lean_dec(x_20); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_299 = !lean_is_exclusive(x_26); +x_299 = !lean_is_exclusive(x_22); if (x_299 == 0) { -return x_26; +return x_22; } else { lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_300 = lean_ctor_get(x_26, 0); -x_301 = lean_ctor_get(x_26, 1); +x_300 = lean_ctor_get(x_22, 0); +x_301 = lean_ctor_get(x_22, 1); lean_inc(x_301); lean_inc(x_300); -lean_dec(x_26); +lean_dec(x_22); x_302 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_302, 0, x_300); lean_ctor_set(x_302, 1, x_301); @@ -6464,34 +6478,6 @@ return x_302; } } } -else -{ -uint8_t x_303; -lean_dec(x_20); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_303 = !lean_is_exclusive(x_22); -if (x_303 == 0) -{ -return x_22; -} -else -{ -lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_304 = lean_ctor_get(x_22, 0); -x_305 = lean_ctor_get(x_22, 1); -lean_inc(x_305); -lean_inc(x_304); -lean_dec(x_22); -x_306 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_306, 0, x_304); -lean_ctor_set(x_306, 1, x_305); -return x_306; -} -} -} } } LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -6652,8 +6638,6 @@ l_Lean_Meta_reportDiag___closed__14 = _init_l_Lean_Meta_reportDiag___closed__14( lean_mark_persistent(l_Lean_Meta_reportDiag___closed__14); l_Lean_Meta_reportDiag___closed__15 = _init_l_Lean_Meta_reportDiag___closed__15(); lean_mark_persistent(l_Lean_Meta_reportDiag___closed__15); -l_Lean_Meta_reportDiag___closed__16 = _init_l_Lean_Meta_reportDiag___closed__16(); -lean_mark_persistent(l_Lean_Meta_reportDiag___closed__16); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/DiscrTree.c b/stage0/stdlib/Lean/Meta/DiscrTree.c index 9d5c6ca09ab1..abcc4cf3f1a6 100644 --- a/stage0/stdlib/Lean/Meta/DiscrTree.c +++ b/stage0/stdlib/Lean/Meta/DiscrTree.c @@ -189,7 +189,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_DiscrTree_keysAs lean_object* l_Std_Format_join(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_DiscrTree_toArray___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_instLTKey; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_isOffset(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getAllValuesForKey___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_DiscrTree_getUnify___spec__2(lean_object*); @@ -1798,8 +1798,7 @@ x_9 = lean_ctor_get(x_6, 1); x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); -lean_inc(x_1); -x_11 = lean_environment_find(x_10, x_1); +x_11 = l_Lean_Environment_find_x3f(x_10, x_1); if (lean_obj_tag(x_11) == 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; @@ -1839,8 +1838,7 @@ lean_dec(x_6); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -lean_inc(x_1); -x_23 = lean_environment_find(x_22, x_1); +x_23 = l_Lean_Environment_find_x3f(x_22, x_1); if (lean_obj_tag(x_23) == 0) { uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -12932,7 +12930,6 @@ lean_dec(x_68); if (lean_obj_tag(x_69) == 0) { lean_object* x_70; lean_object* x_71; uint8_t x_72; -lean_inc(x_50); x_70 = l_Lean_isRec___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___spec__1(x_50, x_4, x_5, x_6, x_7, x_67); x_71 = lean_ctor_get(x_70, 0); lean_inc(x_71); @@ -13346,6 +13343,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_7; } } diff --git a/stage0/stdlib/Lean/Meta/Eqns.c b/stage0/stdlib/Lean/Meta/Eqns.c index cc7ab35d3ab1..d9d6d6efc1bc 100644 --- a/stage0/stdlib/Lean/Meta/Eqns.c +++ b/stage0/stdlib/Lean/Meta/Eqns.c @@ -69,7 +69,7 @@ lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Eqns_0__Lean_Meta_registerEqnThms___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Eqns_0__Lean_Meta_getEqnsFor_x3fCore___closed__5; static lean_object* l___private_Lean_Meta_Eqns_0__Lean_Meta_getEqnsFor_x3fCore___lambda__2___closed__2; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_registerGetUnfoldEqnFn(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_267____lambda__1___boxed(lean_object*, lean_object*); @@ -86,10 +86,8 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_2302____lambd LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_350_(lean_object*); static lean_object* l_Lean_Meta_eqnThmSuffixBasePrefix___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_backward_eqns_deepRecursiveSplit; -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_98____closed__2; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_getUnfoldEqnFor_x3f___spec__1___closed__2; lean_object* l_Lean_initializing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_5_(lean_object*); @@ -120,6 +118,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Eqns LEAN_EXPORT lean_object* l___private_Lean_Meta_Eqns_0__Lean_Meta_shouldGenerateEqnThms(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_eqn1ThmSuffix___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Eqns_0__Lean_Meta_alreadyGenerated_x3f_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Meta_eqnAffectingOptions___closed__2; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_2302____lambda__2___closed__8; static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_getUnfoldEqnFor_x3f___spec__1___closed__1; @@ -259,6 +258,7 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Eqns_0__ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_isEqnThm_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_Eqns_0__Lean_Meta_registerEqnThms___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_getUnfoldEqnFor_x3f___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_98____closed__1; @@ -267,6 +267,7 @@ lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Meta_registerGetEqnsFn___closed__2; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); lean_object* l_Lean_mkTagDeclarationExtension(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Eqns_0__Lean_Meta_shouldGenerateEqnThms___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwReservedNameNotAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_isEqnThm_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_unfoldThmSuffix___closed__1; @@ -1202,7 +1203,6 @@ return x_10; else { uint8_t x_11; -lean_inc(x_3); lean_inc(x_1); x_11 = l_Lean_Environment_isSafeDefinition(x_1, x_3); if (x_11 == 0) @@ -1236,7 +1236,6 @@ else { uint8_t x_16; lean_dec(x_4); -lean_inc(x_3); lean_inc(x_1); x_16 = l_Lean_Environment_isSafeDefinition(x_1, x_3); if (x_16 == 0) @@ -1270,7 +1269,6 @@ else { uint8_t x_21; lean_dec(x_4); -lean_inc(x_3); lean_inc(x_1); x_21 = l_Lean_Environment_isSafeDefinition(x_1, x_3); if (x_21 == 0) @@ -1550,7 +1548,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -x_12 = lean_environment_find(x_11, x_1); +x_12 = l_Lean_Environment_find_x3f(x_11, x_1); if (lean_obj_tag(x_12) == 0) { uint8_t x_13; lean_object* x_14; @@ -1691,7 +1689,7 @@ lean_dec(x_7); x_42 = lean_ctor_get(x_40, 0); lean_inc(x_42); lean_dec(x_40); -x_43 = lean_environment_find(x_42, x_1); +x_43 = l_Lean_Environment_find_x3f(x_42, x_1); if (lean_obj_tag(x_43) == 0) { uint8_t x_44; lean_object* x_45; lean_object* x_46; @@ -1837,6 +1835,15 @@ lean_dec(x_1); return x_7; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Eqns_0__Lean_Meta_shouldGenerateEqnThms___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Lean_Meta_Eqns_0__Lean_Meta_shouldGenerateEqnThms(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_1); +return x_7; +} +} static lean_object* _init_l_Lean_Meta_instInhabitedEqnsExtState___closed__1() { _start: { @@ -2475,8 +2482,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; @@ -2540,8 +2546,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; @@ -4319,7 +4324,6 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_1); x_21 = l___private_Lean_Meta_Eqns_0__Lean_Meta_shouldGenerateEqnThms(x_1, x_2, x_3, x_4, x_5, x_20); if (lean_obj_tag(x_21) == 0) { @@ -4593,7 +4597,6 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_1); x_79 = l___private_Lean_Meta_Eqns_0__Lean_Meta_shouldGenerateEqnThms(x_1, x_2, x_3, x_4, x_5, x_78); if (lean_obj_tag(x_79) == 0) { @@ -5115,7 +5118,7 @@ x_14 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler_ x_15 = lean_ctor_get(x_11, 0); lean_inc(x_15); lean_dec(x_11); -x_16 = l_Lean_Kernel_isDiagnosticsEnabled(x_15); +x_16 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_15); lean_dec(x_15); if (x_16 == 0) { @@ -5170,7 +5173,7 @@ x_22 = lean_ctor_get(x_18, 1); x_23 = lean_ctor_get(x_20, 0); x_24 = lean_ctor_get(x_20, 4); lean_dec(x_24); -x_25 = l_Lean_Kernel_enableDiag(x_23, x_14); +x_25 = l_Lean_Kernel_Environment_enableDiag(x_23, x_14); lean_inc(x_2); lean_ctor_set(x_18, 1, x_2); lean_ctor_set(x_18, 0, x_2); @@ -5203,7 +5206,7 @@ lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); lean_dec(x_20); -x_38 = l_Lean_Kernel_enableDiag(x_31, x_14); +x_38 = l_Lean_Kernel_Environment_enableDiag(x_31, x_14); lean_inc(x_2); lean_ctor_set(x_18, 1, x_2); lean_ctor_set(x_18, 0, x_2); @@ -5261,7 +5264,7 @@ if (lean_is_exclusive(x_44)) { lean_dec_ref(x_44); x_53 = lean_box(0); } -x_54 = l_Lean_Kernel_enableDiag(x_46, x_14); +x_54 = l_Lean_Kernel_Environment_enableDiag(x_46, x_14); lean_inc(x_2); x_55 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_55, 0, x_2); @@ -6200,7 +6203,6 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_1); x_10 = l___private_Lean_Meta_Eqns_0__Lean_Meta_shouldGenerateEqnThms(x_1, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_10) == 0) { @@ -6939,7 +6941,6 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_5); x_12 = l_Lean_Environment_isSafeDefinition(x_11, x_5); if (x_12 == 0) { @@ -6973,7 +6974,6 @@ lean_dec(x_7); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -lean_inc(x_5); x_20 = l_Lean_Environment_isSafeDefinition(x_19, x_5); if (x_20 == 0) { diff --git a/stage0/stdlib/Lean/Meta/ExprDefEq.c b/stage0/stdlib/Lean/Meta/ExprDefEq.c index 1dbc1db95d06..829f9d2d7c0a 100644 --- a/stage0/stdlib/Lean/Meta/ExprDefEq.c +++ b/stage0/stdlib/Lean/Meta/ExprDefEq.c @@ -245,7 +245,7 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___privat LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_checkFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOnFailure___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_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_CheckAssignment_checkApp___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*); @@ -3746,6 +3746,7 @@ lean_dec(x_50); x_53 = lean_ctor_get(x_3, 1); lean_inc(x_53); x_54 = l_Lean_isStructureLike(x_52, x_53); +lean_dec(x_53); if (x_54 == 0) { lean_object* x_55; lean_object* x_56; uint8_t x_57; @@ -4071,6 +4072,7 @@ lean_dec(x_127); x_130 = lean_ctor_get(x_3, 1); lean_inc(x_130); x_131 = l_Lean_isStructureLike(x_129, x_130); +lean_dec(x_130); if (x_131 == 0) { lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; @@ -4480,7 +4482,8 @@ x_14 = lean_ctor_get(x_11, 1); x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); lean_dec(x_13); -x_16 = lean_environment_find(x_15, x_9); +x_16 = l_Lean_Environment_find_x3f(x_15, x_9); +lean_dec(x_9); if (lean_obj_tag(x_16) == 0) { uint8_t x_17; lean_object* x_18; @@ -4575,7 +4578,8 @@ x_39 = lean_ctor_get(x_36, 1); x_40 = lean_ctor_get(x_38, 0); lean_inc(x_40); lean_dec(x_38); -x_41 = lean_environment_find(x_40, x_35); +x_41 = l_Lean_Environment_find_x3f(x_40, x_35); +lean_dec(x_35); if (lean_obj_tag(x_41) == 0) { lean_object* x_42; @@ -4627,7 +4631,8 @@ lean_dec(x_36); x_49 = lean_ctor_get(x_47, 0); lean_inc(x_49); lean_dec(x_47); -x_50 = lean_environment_find(x_49, x_35); +x_50 = l_Lean_Environment_find_x3f(x_49, x_35); +lean_dec(x_35); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; @@ -4707,7 +4712,8 @@ lean_dec(x_11); x_62 = lean_ctor_get(x_60, 0); lean_inc(x_62); lean_dec(x_60); -x_63 = lean_environment_find(x_62, x_9); +x_63 = l_Lean_Environment_find_x3f(x_62, x_9); +lean_dec(x_9); if (lean_obj_tag(x_63) == 0) { uint8_t x_64; lean_object* x_65; lean_object* x_66; @@ -4805,7 +4811,8 @@ if (lean_is_exclusive(x_81)) { x_85 = lean_ctor_get(x_82, 0); lean_inc(x_85); lean_dec(x_82); -x_86 = lean_environment_find(x_85, x_80); +x_86 = l_Lean_Environment_find_x3f(x_85, x_80); +lean_dec(x_80); if (lean_obj_tag(x_86) == 0) { lean_object* x_87; @@ -96642,7 +96649,6 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_1); x_15 = l_Lean_getStructureLikeCtor_x3f(x_14, x_1); if (lean_obj_tag(x_15) == 0) { @@ -96653,7 +96659,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_16 = 0; x_17 = lean_box(x_16); lean_ctor_set(x_10, 0, x_17); @@ -96680,7 +96685,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_22 = 0; x_23 = lean_box(x_22); lean_ctor_set(x_10, 0, x_23); @@ -96692,7 +96696,6 @@ lean_object* x_24; lean_object* x_25; lean_free_object(x_10); x_24 = lean_box(0); x_25 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___lambda__3(x_2, x_18, x_3, x_1, x_24, x_5, x_6, x_7, x_8, x_13); -lean_dec(x_1); return x_25; } } @@ -96708,7 +96711,6 @@ lean_dec(x_10); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -lean_inc(x_1); x_29 = l_Lean_getStructureLikeCtor_x3f(x_28, x_1); if (lean_obj_tag(x_29) == 0) { @@ -96719,7 +96721,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_30 = 0; x_31 = lean_box(x_30); x_32 = lean_alloc_ctor(0, 2, 0); @@ -96748,7 +96749,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_37 = 0; x_38 = lean_box(x_37); x_39 = lean_alloc_ctor(0, 2, 0); @@ -96761,7 +96761,6 @@ else lean_object* x_40; lean_object* x_41; x_40 = lean_box(0); x_41 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___lambda__3(x_2, x_33, x_3, x_1, x_40, x_5, x_6, x_7, x_8, x_27); -lean_dec(x_1); return x_41; } } @@ -96790,6 +96789,7 @@ lean_object* x_15; lean_object* x_16; lean_free_object(x_9); x_15 = lean_box(0); x_16 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___lambda__4(x_1, x_2, x_3, x_15, x_4, x_5, x_6, x_7, x_12); +lean_dec(x_1); return x_16; } else @@ -96826,6 +96826,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; x_23 = lean_box(0); x_24 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___lambda__4(x_1, x_2, x_3, x_23, x_4, x_5, x_6, x_7, x_20); +lean_dec(x_1); return x_24; } else @@ -96884,6 +96885,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqPro lean_object* x_10; x_10 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_4); +lean_dec(x_1); return x_10; } } @@ -100012,7 +100014,8 @@ x_20 = lean_ctor_get(x_17, 1); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -x_22 = lean_environment_find(x_21, x_16); +x_22 = l_Lean_Environment_find_x3f(x_21, x_16); +lean_dec(x_16); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; @@ -100485,7 +100488,8 @@ lean_dec(x_17); x_114 = lean_ctor_get(x_112, 0); lean_inc(x_114); lean_dec(x_112); -x_115 = lean_environment_find(x_114, x_16); +x_115 = l_Lean_Environment_find_x3f(x_114, x_16); +lean_dec(x_16); if (lean_obj_tag(x_115) == 0) { uint8_t x_116; lean_object* x_117; lean_object* x_118; @@ -100885,7 +100889,8 @@ if (lean_is_exclusive(x_186)) { x_190 = lean_ctor_get(x_187, 0); lean_inc(x_190); lean_dec(x_187); -x_191 = lean_environment_find(x_190, x_185); +x_191 = l_Lean_Environment_find_x3f(x_190, x_185); +lean_dec(x_185); if (lean_obj_tag(x_191) == 0) { uint8_t x_192; lean_object* x_193; lean_object* x_194; diff --git a/stage0/stdlib/Lean/Meta/GetUnfoldableConst.c b/stage0/stdlib/Lean/Meta/GetUnfoldableConst.c index 74fe94f25055..cbad785ef3f1 100644 --- a/stage0/stdlib/Lean/Meta/GetUnfoldableConst.c +++ b/stage0/stdlib/Lean/Meta/GetUnfoldableConst.c @@ -18,7 +18,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_GetUnfoldableConst_0__Lean_Meta_c LEAN_EXPORT lean_object* l_Lean_isIrreducible___at___private_Lean_Meta_GetUnfoldableConst_0__Lean_Meta_canUnfoldDefault___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_is_instance(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_name(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_canUnfold___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_get_reducibility_status(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_GetUnfoldableConst_0__Lean_Meta_canUnfoldDefault___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -538,8 +538,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_1); -x_12 = lean_environment_find(x_11, x_1); +x_12 = l_Lean_Environment_find_x3f(x_11, x_1); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; @@ -803,8 +802,7 @@ lean_dec(x_7); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -lean_inc(x_1); -x_55 = lean_environment_find(x_54, x_1); +x_55 = l_Lean_Environment_find_x3f(x_54, x_1); if (lean_obj_tag(x_55) == 0) { lean_object* x_56; @@ -976,7 +974,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -x_12 = lean_environment_find(x_11, x_1); +x_12 = l_Lean_Environment_find_x3f(x_11, x_1); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; @@ -1239,7 +1237,7 @@ lean_dec(x_7); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_find(x_54, x_1); +x_55 = l_Lean_Environment_find_x3f(x_54, x_1); if (lean_obj_tag(x_55) == 0) { lean_object* x_56; lean_object* x_57; @@ -1396,6 +1394,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getUnfoldableConstNoEx_x3f___boxed(lean_obj lean_object* x_7; x_7 = l_Lean_Meta_getUnfoldableConstNoEx_x3f(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); +lean_dec(x_1); return x_7; } } diff --git a/stage0/stdlib/Lean/Meta/IndPredBelow.c b/stage0/stdlib/Lean/Meta/IndPredBelow.c index 174f7baf6ce7..2c95aa8b2735 100644 --- a/stage0/stdlib/Lean/Meta/IndPredBelow.c +++ b/stage0/stdlib/Lean/Meta/IndPredBelow.c @@ -23583,6 +23583,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); x_36 = l_Lean_Meta_isInductivePredicate(x_30, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_30); if (lean_obj_tag(x_36) == 0) { lean_object* x_37; uint8_t x_38; @@ -25088,7 +25089,6 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_1); x_7 = l_Lean_Meta_isInductivePredicate(x_1, x_2, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_7) == 0) { diff --git a/stage0/stdlib/Lean/Meta/InferType.c b/stage0/stdlib/Lean/Meta/InferType.c index a1772587d1e3..7357e5527c1e 100644 --- a/stage0/stdlib/Lean/Meta/InferType.c +++ b/stage0/stdlib/Lean/Meta/InferType.c @@ -85,7 +85,7 @@ lean_object* l_Lean_Meta_forallTelescope___at_Lean_Meta_mapForallTelescope_x27__ static lean_object* l_Lean_Expr_instantiateBetaRevRange___closed__1; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwUnknownMVar___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_arrowDomainsN___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwTypeExcepted___rarg___closed__2; @@ -5109,7 +5109,8 @@ x_21 = lean_ctor_get(x_18, 1); x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); lean_dec(x_20); -x_23 = lean_environment_find(x_22, x_16); +x_23 = l_Lean_Environment_find_x3f(x_22, x_16); +lean_dec(x_16); if (lean_obj_tag(x_23) == 0) { lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; @@ -5671,7 +5672,8 @@ lean_dec(x_18); x_172 = lean_ctor_get(x_170, 0); lean_inc(x_172); lean_dec(x_170); -x_173 = lean_environment_find(x_172, x_16); +x_173 = l_Lean_Environment_find_x3f(x_172, x_16); +lean_dec(x_16); if (lean_obj_tag(x_173) == 0) { lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; diff --git a/stage0/stdlib/Lean/Meta/Injective.c b/stage0/stdlib/Lean/Meta/Injective.c index dd8dc9b3e61b..48e6a1fccb7e 100644 --- a/stage0/stdlib/Lean/Meta/Injective.c +++ b/stage0/stdlib/Lean/Meta/Injective.c @@ -65,7 +65,7 @@ lean_object* l_Lean_Meta_forallTelescope___at_Lean_Meta_mapForallTelescope_x27__ LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1973____closed__3; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremType_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_refl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5581,8 +5581,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_1); -x_12 = lean_environment_find(x_11, x_1); +x_12 = l_Lean_Environment_find_x3f(x_11, x_1); if (lean_obj_tag(x_12) == 0) { 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; @@ -5622,8 +5621,7 @@ lean_dec(x_7); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -lean_inc(x_1); -x_24 = lean_environment_find(x_23, x_1); +x_24 = l_Lean_Environment_find_x3f(x_23, x_1); if (lean_obj_tag(x_24) == 0) { uint8_t 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; @@ -8372,7 +8370,6 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_1); x_12 = l_Lean_Meta_isInductivePredicate(x_1, x_2, x_3, x_4, x_5, x_9); if (lean_obj_tag(x_12) == 0) { diff --git a/stage0/stdlib/Lean/Meta/Instances.c b/stage0/stdlib/Lean/Meta/Instances.c index 30102f6393b3..72df63ef856a 100644 --- a/stage0/stdlib/Lean/Meta/Instances.c +++ b/stage0/stdlib/Lean/Meta/Instances.c @@ -99,7 +99,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_I static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_3714____closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedInstanceEntry; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__7___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_4659____closed__1; static lean_object* l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -15407,8 +15407,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); 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; @@ -15457,8 +15456,7 @@ lean_dec(x_8); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_1); -x_27 = lean_environment_find(x_26, x_1); +x_27 = l_Lean_Environment_find_x3f(x_26, x_1); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; diff --git a/stage0/stdlib/Lean/Meta/LazyDiscrTree.c b/stage0/stdlib/Lean/Meta/LazyDiscrTree.c index 07e2d479950b..948a3a99c3b9 100644 --- a/stage0/stdlib/Lean/Meta/LazyDiscrTree.c +++ b/stage0/stdlib/Lean/Meta/LazyDiscrTree.c @@ -185,11 +185,9 @@ static lean_object* l_Lean_Meta_LazyDiscrTree_instInhabitedTrie___closed__1; static uint64_t l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_getStarResult___rarg___closed__5; static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData___rarg___lambda__2___closed__12; static lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree___at_Lean_Meta_LazyDiscrTree_findImportMatches___spec__1___rarg___closed__2; -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_createImportedEnvironmentSeq_go(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_LazyDiscrTree_PreDiscrTree_append___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_MatchResult_size___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_setTrie(lean_object*); static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData___rarg___lambda__2___closed__4; static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData___rarg___lambda__2___closed__13; @@ -259,6 +257,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree_go___ static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_reprKey____x40_Lean_Meta_LazyDiscrTree___hyg_355____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_getKeyArgs(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_ImportData_new(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_logAt___at_Lean_Meta_LazyDiscrTree_findImportMatches___spec__7___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_LazyDiscrTree_createLocalPreDiscrTree___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_reprKey____x40_Lean_Meta_LazyDiscrTree___hyg_355____closed__21; @@ -571,6 +570,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_LazyDiscrTree_Mat LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addLazyEntryToTrie___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_InitResults_instAppendInitResults(lean_object*); lean_object* lean_array_get_size(lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_targetPath___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_evalLazyEntries___spec__1___rarg(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree_go___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*); @@ -3597,7 +3597,6 @@ lean_dec(x_68); if (lean_obj_tag(x_69) == 0) { lean_object* x_70; lean_object* x_71; uint8_t x_72; -lean_inc(x_50); x_70 = l_Lean_isRec___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___spec__1(x_50, x_4, x_5, x_6, x_7, x_67); x_71 = lean_ctor_get(x_70, 0); lean_inc(x_71); @@ -16786,7 +16785,7 @@ lean_dec(x_94); x_97 = lean_ctor_get(x_95, 0); lean_inc(x_97); lean_dec(x_95); -x_98 = l_Lean_Kernel_isDiagnosticsEnabled(x_97); +x_98 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_97); lean_dec(x_97); if (x_98 == 0) { @@ -16840,7 +16839,7 @@ lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; x_104 = lean_ctor_get(x_101, 0); x_105 = lean_ctor_get(x_101, 4); lean_dec(x_105); -x_106 = l_Lean_Kernel_enableDiag(x_104, x_93); +x_106 = l_Lean_Kernel_Environment_enableDiag(x_104, x_93); x_107 = l_Lean_Meta_LazyDiscrTree_Cache_empty___closed__3; lean_ctor_set(x_101, 4, x_107); lean_ctor_set(x_101, 0, x_106); @@ -16929,7 +16928,7 @@ lean_inc(x_129); lean_inc(x_128); lean_inc(x_127); lean_dec(x_101); -x_134 = l_Lean_Kernel_enableDiag(x_127, x_93); +x_134 = l_Lean_Kernel_Environment_enableDiag(x_127, x_93); x_135 = l_Lean_Meta_LazyDiscrTree_Cache_empty___closed__3; x_136 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_136, 0, x_134); @@ -17851,7 +17850,7 @@ return x_11; else { lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_12 = lean_ctor_get(x_2, 4); +x_12 = lean_ctor_get(x_2, 5); lean_inc(x_12); x_13 = lean_ctor_get(x_12, 3); lean_inc(x_13); @@ -18732,7 +18731,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createLocalPreDiscrTree___rar _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_ctor_get(x_3, 4); +x_7 = lean_ctor_get(x_3, 5); lean_inc(x_7); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); @@ -18755,7 +18754,7 @@ lean_closure_set(x_13, 2, x_8); lean_closure_set(x_13, 3, x_4); lean_closure_set(x_13, 4, x_11); lean_closure_set(x_13, 5, x_5); -x_14 = lean_ctor_get(x_3, 1); +x_14 = lean_ctor_get(x_3, 0); lean_inc(x_14); lean_dec(x_3); x_15 = lean_ctor_get(x_14, 1); @@ -19139,7 +19138,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree_go___ _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_ctor_get(x_4, 4); +x_13 = lean_ctor_get(x_4, 5); lean_inc(x_13); x_14 = lean_ctor_get(x_13, 4); lean_inc(x_14); @@ -19596,7 +19595,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree___rar _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_11 = lean_ctor_get(x_8, 4); +x_11 = lean_ctor_get(x_8, 5); lean_inc(x_11); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); @@ -19780,7 +19779,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree_go___ _start: { lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_2, 4); +x_16 = lean_ctor_get(x_2, 5); lean_inc(x_16); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); @@ -21178,7 +21177,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree___at_ _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_11 = lean_ctor_get(x_3, 4); +x_11 = lean_ctor_get(x_3, 5); lean_inc(x_11); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); diff --git a/stage0/stdlib/Lean/Meta/Match/Basic.c b/stage0/stdlib/Lean/Meta/Match/Basic.c index caa9167901fa..c4c5f79fb684 100644 --- a/stage0/stdlib/Lean/Meta/Match/Basic.c +++ b/stage0/stdlib/Lean/Meta/Match/Basic.c @@ -62,7 +62,7 @@ lean_object* l_Lean_Expr_bvar___override(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Match_Example_applyFVarSubst___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_FVarSubst_get(lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_instantiatePatternMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_arrayLit_x3f(lean_object*); static lean_object* l_Lean_Meta_Match_Pattern_toMessageData___closed__2; @@ -7340,7 +7340,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_find(x_29, x_23); +x_30 = l_Lean_Environment_find_x3f(x_29, x_23); +lean_dec(x_23); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -7478,7 +7479,8 @@ lean_dec(x_25); x_68 = lean_ctor_get(x_66, 0); lean_inc(x_68); lean_dec(x_66); -x_69 = lean_environment_find(x_68, x_23); +x_69 = l_Lean_Environment_find_x3f(x_68, x_23); +lean_dec(x_23); if (lean_obj_tag(x_69) == 0) { lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; @@ -7733,7 +7735,8 @@ if (lean_is_exclusive(x_130)) { x_134 = lean_ctor_get(x_131, 0); lean_inc(x_134); lean_dec(x_131); -x_135 = lean_environment_find(x_134, x_128); +x_135 = l_Lean_Environment_find_x3f(x_134, x_128); +lean_dec(x_128); if (lean_obj_tag(x_135) == 0) { lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; diff --git a/stage0/stdlib/Lean/Meta/Match/MatchEqs.c b/stage0/stdlib/Lean/Meta/Match/MatchEqs.c index 33b7734d0c14..e94cd8cf6e31 100644 --- a/stage0/stdlib/Lean/Meta/Match/MatchEqs.c +++ b/stage0/stdlib/Lean/Meta/Match/MatchEqs.c @@ -190,7 +190,7 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_proveCondEqThm___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_injectionAny(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___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_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_Meta_Match_proveCondEqThm_go___lambda__2___closed__9; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_substSomeVar___lambda__1___closed__1; @@ -5258,7 +5258,8 @@ x_26 = lean_ctor_get(x_23, 1); x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); -x_28 = lean_environment_find(x_27, x_8); +x_28 = l_Lean_Environment_find_x3f(x_27, x_8); +lean_dec(x_8); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; @@ -5326,7 +5327,8 @@ lean_dec(x_23); x_41 = lean_ctor_get(x_39, 0); lean_inc(x_41); lean_dec(x_39); -x_42 = lean_environment_find(x_41, x_8); +x_42 = l_Lean_Environment_find_x3f(x_41, x_8); +lean_dec(x_8); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; @@ -48608,6 +48610,7 @@ lean_inc(x_17); lean_dec(x_15); lean_inc(x_1); x_18 = l_Lean_mkPrivateName(x_17, x_1); +lean_dec(x_17); x_19 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___lambda__5___closed__2; lean_inc(x_18); x_20 = l_Lean_Name_append(x_18, x_19); @@ -48838,6 +48841,7 @@ lean_inc(x_72); lean_dec(x_70); lean_inc(x_1); x_73 = l_Lean_mkPrivateName(x_72, x_1); +lean_dec(x_72); x_74 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___lambda__5___closed__2; lean_inc(x_73); x_75 = l_Lean_Name_append(x_73, x_74); @@ -49088,6 +49092,7 @@ lean_inc(x_139); lean_dec(x_136); lean_inc(x_1); x_140 = l_Lean_mkPrivateName(x_139, x_1); +lean_dec(x_139); x_141 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___lambda__5___closed__2; lean_inc(x_140); x_142 = l_Lean_Name_append(x_140, x_141); @@ -49382,6 +49387,7 @@ lean_inc(x_218); lean_dec(x_215); lean_inc(x_1); x_219 = l_Lean_mkPrivateName(x_218, x_1); +lean_dec(x_218); x_220 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___lambda__5___closed__2; lean_inc(x_219); x_221 = l_Lean_Name_append(x_219, x_220); diff --git a/stage0/stdlib/Lean/Meta/SizeOf.c b/stage0/stdlib/Lean/Meta/SizeOf.c index 1b87662a8c16..14b09c67a5ff 100644 --- a/stage0/stdlib/Lean/Meta/SizeOf.c +++ b/stage0/stdlib/Lean/Meta/SizeOf.c @@ -109,7 +109,7 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_7441____clo static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8191____closed__6; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn___lambda__6___boxed(lean_object**); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_RecursorVal_getFirstMinorIdx(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__8___boxed(lean_object**); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProofStep___closed__2; @@ -7539,7 +7539,8 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -x_14 = lean_environment_find(x_13, x_8); +x_14 = l_Lean_Environment_find_x3f(x_13, x_8); +lean_dec(x_8); if (lean_obj_tag(x_14) == 0) { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -7746,7 +7747,8 @@ lean_dec(x_9); x_75 = lean_ctor_get(x_73, 0); lean_inc(x_75); lean_dec(x_73); -x_76 = lean_environment_find(x_75, x_8); +x_76 = l_Lean_Environment_find_x3f(x_75, x_8); +lean_dec(x_8); if (lean_obj_tag(x_76) == 0) { lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; @@ -8290,7 +8292,8 @@ x_14 = lean_ctor_get(x_11, 1); x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); lean_dec(x_13); -x_16 = lean_environment_find(x_15, x_9); +x_16 = l_Lean_Environment_find_x3f(x_15, x_9); +lean_dec(x_9); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; @@ -8457,7 +8460,8 @@ lean_dec(x_11); x_75 = lean_ctor_get(x_73, 0); lean_inc(x_75); lean_dec(x_73); -x_76 = lean_environment_find(x_75, x_9); +x_76 = l_Lean_Environment_find_x3f(x_75, x_9); +lean_dec(x_9); if (lean_obj_tag(x_76) == 0) { lean_object* x_77; @@ -11251,7 +11255,8 @@ lean_dec(x_33); x_36 = lean_ctor_get(x_34, 0); lean_inc(x_36); lean_dec(x_34); -x_37 = lean_environment_find(x_36, x_32); +x_37 = l_Lean_Environment_find_x3f(x_36, x_32); +lean_dec(x_32); if (lean_obj_tag(x_37) == 0) { lean_object* x_38; @@ -11547,7 +11552,8 @@ lean_dec(x_94); x_97 = lean_ctor_get(x_95, 0); lean_inc(x_97); lean_dec(x_95); -x_98 = lean_environment_find(x_97, x_93); +x_98 = l_Lean_Environment_find_x3f(x_97, x_93); +lean_dec(x_93); if (lean_obj_tag(x_98) == 0) { lean_object* x_99; @@ -12023,8 +12029,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -12064,8 +12069,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; @@ -18318,7 +18322,6 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_1); x_12 = l_Lean_Meta_isInductivePredicate(x_1, x_2, x_3, x_4, x_5, x_9); if (lean_obj_tag(x_12) == 0) { diff --git a/stage0/stdlib/Lean/Meta/Sorry.c b/stage0/stdlib/Lean/Meta/Sorry.c index d7abfefff979..542cda860d47 100644 --- a/stage0/stdlib/Lean/Meta/Sorry.c +++ b/stage0/stdlib/Lean/Meta/Sorry.c @@ -64,7 +64,7 @@ uint8_t l_Lean_Name_hasMacroScopes(lean_object*); static lean_object* l_Lean_Meta_mkSorry___closed__6; static lean_object* l_Lean_Meta_mkLabeledSorry___lambda__1___closed__10; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkLabeledSorry___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_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSorry___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -911,7 +911,8 @@ lean_dec(x_27); x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); +x_31 = l_Lean_Environment_mainModule(x_30); +lean_dec(x_30); lean_inc(x_26); x_32 = l_Lean_FileMap_toPosition(x_26, x_17); lean_inc(x_26); @@ -997,7 +998,8 @@ lean_dec(x_57); x_60 = lean_ctor_get(x_58, 0); lean_inc(x_60); lean_dec(x_58); -x_61 = lean_environment_main_module(x_60); +x_61 = l_Lean_Environment_mainModule(x_60); +lean_dec(x_60); lean_inc(x_56); x_62 = l_Lean_FileMap_toPosition(x_56, x_17); lean_inc(x_56); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Apply.c b/stage0/stdlib/Lean/Meta/Tactic/Apply.c index e8b2616cb4e6..76530a5c0106 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Apply.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Apply.c @@ -59,7 +59,7 @@ lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_MVarId_iffOfEq___closed__1; lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_elem___at_Lean_MVarId_apply___spec__1___boxed(lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_synthAppInstances_step___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getExpectedNumArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -6700,7 +6700,8 @@ lean_dec(x_18); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -x_22 = lean_environment_find(x_21, x_16); +x_22 = l_Lean_Environment_find_x3f(x_21, x_16); +lean_dec(x_16); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; diff --git a/stage0/stdlib/Lean/Meta/Tactic/AuxLemma.c b/stage0/stdlib/Lean/Meta/Tactic/AuxLemma.c index 73e81c909189..ae3688d3eb88 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/AuxLemma.c +++ b/stage0/stdlib/Lean/Meta/Tactic/AuxLemma.c @@ -54,7 +54,7 @@ static lean_object* l_Lean_Meta_instInhabitedAuxLemmas___closed__1; static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_mkAuxLemma___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_mkAuxLemma___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkAuxLemma___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_auxLemmasExt; @@ -984,7 +984,8 @@ lean_dec(x_11); x_14 = l_Lean_Meta_instInhabitedAuxLemmas; x_15 = l_Lean_Meta_mkAuxLemma___closed__1; x_16 = l_Lean_EnvExtension_getState___rarg(x_14, x_15, x_13); -x_17 = lean_environment_main_module(x_13); +x_17 = l_Lean_Environment_mainModule(x_13); +lean_dec(x_13); x_18 = l_Lean_Meta_mkAuxLemma___closed__3; x_19 = l_Lean_Name_append(x_17, x_18); x_20 = lean_ctor_get(x_16, 0); @@ -1529,7 +1530,8 @@ lean_dec(x_163); x_166 = l_Lean_Meta_instInhabitedAuxLemmas; x_167 = l_Lean_Meta_mkAuxLemma___closed__1; x_168 = l_Lean_EnvExtension_getState___rarg(x_166, x_167, x_165); -x_169 = lean_environment_main_module(x_165); +x_169 = l_Lean_Environment_mainModule(x_165); +lean_dec(x_165); x_170 = l_Lean_Meta_mkAuxLemma___closed__3; x_171 = l_Lean_Name_append(x_169, x_170); x_172 = lean_ctor_get(x_168, 0); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Cases.c b/stage0/stdlib/Lean/Meta/Tactic/Cases.c index 662aaea31e43..7e60a2fdccb9 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Cases.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Cases.c @@ -14,28 +14,27 @@ extern "C" { #endif static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___closed__1; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_FVarSubst_apply(lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__3; static lean_object* l_Lean_MVarId_byCasesDec___closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__44___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_casesAnd___closed__3; lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_elimAuxIndices___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__25(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__45___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__40___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__37(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -52,6 +51,7 @@ extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__24___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3; lean_object* l_Lean_MVarId_checkNotAssigned(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); @@ -62,6 +62,7 @@ static lean_object* l_Lean_Meta_Cases_unifyEqs_x3f___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__41___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1; uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__2; @@ -76,8 +77,8 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__31___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__53(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_anyTR_loop___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2; static lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__2; lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -90,29 +91,29 @@ lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__35(lean_object*, lean_object*, lean_object*, size_t, size_t); extern lean_object* l_Lean_casesOnSuffix; LEAN_EXPORT lean_object* l_Lean_MVarId_byCasesDec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__41(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__6; lean_object* l_Lean_Meta_FVarSubst_get(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withNewEqs___rarg___closed__1; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__47___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10; +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_throwNestedTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__8___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withNewEqs_loop___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_elimAuxIndices___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__3(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__1; lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_generalizeTargetsEq___closed__1; @@ -129,31 +130,36 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__57(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_casesRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_elimAuxIndices___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___closed__4; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_casesAnd___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__37___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__3; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1___closed__6; lean_object* l_Lean_AssocList_erase___at_Lean_Meta_FVarSubst_erase___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_toCasesSubgoals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__52___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__55___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__40(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__1; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2; static lean_object* l_Lean_Meta_generalizeTargetsEq___closed__2; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__39(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -161,18 +167,18 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_h LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__36___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__23(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__2; -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__7___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_MVarId_substEqs___closed__1; static lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10; static lean_object* l_Lean_MVarId_byCases___closed__1; static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__10; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__48___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1; +static lean_object* l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2; uint8_t lean_expr_eqv(lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__36(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -182,17 +188,18 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f_ lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isEq(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__54___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__53___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__4; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6; LEAN_EXPORT lean_object* l_Lean_MVarId_byCases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__51(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_casesRec___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__30(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_acyclic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___closed__2; @@ -200,17 +207,16 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_casesRec___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794_(lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__9; lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__7; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Nat_anyTR_loop___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -220,11 +226,10 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Ta LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__8; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1; lean_object* l_Lean_MVarId_induction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__44(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_toCasesSubgoals___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -235,22 +240,24 @@ extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_unifyCasesEqs___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_MVarId_byCases___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop(lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_casesAnd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__16___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_elimAuxIndices___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__31(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__22(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__58___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_MVarId_clear(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__58(lean_object*, size_t, size_t); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6; LEAN_EXPORT lean_object* l_Lean_observing_x3f___at_Lean_MVarId_casesRec___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_byCasesDec___closed__4; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -269,18 +276,22 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_i LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_casesAnd___closed__4; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__16(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_casesRec___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3; lean_object* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_getInductiveUniverseAndParams___closed__1; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7; static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__5; static lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__33(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_MVarId_substEqs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withNewEqs_loop___rarg___closed__1; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_MVarId_cases(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__15(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_substEqs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -293,7 +304,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__43(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_LocalDecl_userName(lean_object*); lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11; lean_object* l_Lean_Meta_exactlyOne(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_casesRec___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___closed__1; @@ -301,13 +312,16 @@ lean_object* l_Lean_Meta_introNCore(lean_object*, lean_object*, lean_object*, ui LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__34___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__32___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__34(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12; lean_object* l_Lean_Meta_ensureAtMostOne(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__32(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_casesRec___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__39___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f___spec__1___closed__3; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); @@ -324,7 +338,6 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Cases_unifyEqs_x3f_ lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__23___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__20(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); @@ -338,9 +351,9 @@ lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_byCasesDec___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -352,13 +365,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases(lean_object*, lean_object*, lea static lean_object* l_Lean_MVarId_byCases___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__46___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8; lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__18(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* lean_array_mk(lean_object*); lean_object* l_Lean_Meta_saturate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -367,14 +376,15 @@ lean_object* l_Lean_mkEM(lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__6; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__48(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__55(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_MVarId_casesRec___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_toCasesSubgoals___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_MVarId_casesRec___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_MVarId_casesRec___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_unifyCasesEqs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -385,23 +395,21 @@ lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4; lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___spec__1(lean_object*); +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__49(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Cases_cases___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__57___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_generalizeIndices___lambda__1___closed__2; LEAN_EXPORT uint8_t l_Nat_anyTR_loop___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_casesAnd___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop(lean_object*); lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCasesSubgoal___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_casesAnd___closed__2; @@ -416,10 +424,11 @@ static lean_object* l_Lean_MVarId_byCases___closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Cases_unifyEqs_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_generalizeIndices___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_MVarId_casesRec___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_inductionCasesOn___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -434,25 +443,21 @@ lean_object* l_Lean_Expr_mvarId_x21(lean_object*); static lean_object* l_Lean_MVarId_byCases___closed__4; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__17(lean_object*, lean_object*, lean_object*, size_t, size_t); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__19(lean_object*, lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__51___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4; static lean_object* l_Lean_MVarId_casesAnd___closed__1; -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__47(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7; LEAN_EXPORT lean_object* l_Nat_anyTR_loop___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___rarg___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_generalizeTargetsEq___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_unifyCasesEqs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2; +static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_throwInductiveTypeExpected___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -627,7 +632,8 @@ x_16 = lean_ctor_get(x_13, 1); x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -x_18 = lean_environment_find(x_17, x_11); +x_18 = l_Lean_Environment_find_x3f(x_17, x_11); +lean_dec(x_11); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; @@ -703,7 +709,8 @@ lean_dec(x_13); x_35 = lean_ctor_get(x_33, 0); lean_inc(x_35); lean_dec(x_33); -x_36 = lean_environment_find(x_35, x_11); +x_36 = l_Lean_Environment_find_x3f(x_35, x_11); +lean_dec(x_11); if (lean_obj_tag(x_36) == 0) { lean_object* x_37; @@ -1162,7 +1169,7 @@ return x_79; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; @@ -1170,11 +1177,11 @@ x_14 = lean_unsigned_to_nat(1u); x_15 = lean_nat_add(x_1, x_14); x_16 = lean_array_push(x_2, x_8); x_17 = lean_array_push(x_3, x_4); -x_18 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg(x_5, x_6, x_7, x_15, x_16, x_17, x_9, x_10, x_11, x_12, x_13); +x_18 = l_Lean_Meta_withNewEqs_loop___rarg(x_5, x_6, x_7, x_15, x_16, x_17, x_9, x_10, x_11, x_12, x_13); return x_18; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_withNewEqs_loop___rarg___closed__1() { _start: { lean_object* x_1; @@ -1182,17 +1189,17 @@ x_1 = lean_mk_string_unchecked("h", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2() { +static lean_object* _init_l_Lean_Meta_withNewEqs_loop___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1; +x_2 = l_Lean_Meta_withNewEqs_loop___rarg___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; uint8_t x_13; @@ -1238,7 +1245,7 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_21, 1); lean_inc(x_24); lean_dec(x_21); -x_25 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed), 13, 7); +x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed), 13, 7); lean_closure_set(x_25, 0, x_4); lean_closure_set(x_25, 1, x_5); lean_closure_set(x_25, 2, x_6); @@ -1246,7 +1253,7 @@ lean_closure_set(x_25, 3, x_24); lean_closure_set(x_25, 4, x_1); lean_closure_set(x_25, 5, x_2); lean_closure_set(x_25, 6, x_3); -x_26 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2; +x_26 = l_Lean_Meta_withNewEqs_loop___rarg___closed__2; x_27 = 0; x_28 = 0; x_29 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_26, x_27, x_23, x_25, x_28, x_7, x_8, x_9, x_10, x_22); @@ -1307,7 +1314,7 @@ lean_inc(x_38); x_39 = lean_ctor_get(x_36, 1); lean_inc(x_39); lean_dec(x_36); -x_40 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed), 13, 7); +x_40 = lean_alloc_closure((void*)(l_Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed), 13, 7); lean_closure_set(x_40, 0, x_4); lean_closure_set(x_40, 1, x_5); lean_closure_set(x_40, 2, x_6); @@ -1315,7 +1322,7 @@ lean_closure_set(x_40, 3, x_39); lean_closure_set(x_40, 4, x_1); lean_closure_set(x_40, 5, x_2); lean_closure_set(x_40, 6, x_3); -x_41 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2; +x_41 = l_Lean_Meta_withNewEqs_loop___rarg___closed__2; x_42 = 0; x_43 = 0; x_44 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_41, x_42, x_38, x_40, x_43, x_7, x_8, x_9, x_10, x_37); @@ -1357,24 +1364,24 @@ return x_48; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg), 11, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewEqs_loop___rarg), 11, 0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs_loop___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_14 = l_Lean_Meta_withNewEqs_loop___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_1); return x_14; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_withNewEqs___rarg___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -1383,21 +1390,21 @@ x_2 = lean_array_mk(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = lean_unsigned_to_nat(0u); -x_10 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg(x_1, x_2, x_3, x_9, x_10, x_10, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_Meta_withNewEqs___rarg___closed__1; +x_11 = l_Lean_Meta_withNewEqs_loop___rarg(x_1, x_2, x_3, x_9, x_10, x_10, x_4, x_5, x_6, x_7, x_8); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewEqs(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg), 8, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewEqs___rarg), 8, 0); return x_2; } } @@ -1539,7 +1546,7 @@ lean_inc(x_2); x_10 = lean_alloc_closure((void*)(l_Lean_Meta_generalizeTargetsEq___lambda__1___boxed), 9, 2); lean_closure_set(x_10, 0, x_1); lean_closure_set(x_10, 1, x_2); -x_11 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg(x_3, x_2, x_10, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_withNewEqs___rarg(x_3, x_2, x_10, x_5, x_6, x_7, x_8, x_9); return x_11; } } @@ -1892,7 +1899,7 @@ lean_dec(x_4); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; lean_object* x_17; @@ -2261,52 +2268,50 @@ return x_96; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_15; lean_object* x_16; -x_15 = l_Lean_LocalDecl_toExpr(x_1); +lean_object* x_15; lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_2); -lean_inc(x_15); -x_16 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof(x_15, x_2, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_16) == 0) +lean_inc(x_1); +x_15 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; -x_17 = lean_ctor_get(x_16, 0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_17, 0); +x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_array_push(x_9, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1___boxed), 15, 9); -lean_closure_set(x_22, 0, x_8); -lean_closure_set(x_22, 1, x_3); -lean_closure_set(x_22, 2, x_2); -lean_closure_set(x_22, 3, x_4); -lean_closure_set(x_22, 4, x_5); -lean_closure_set(x_22, 5, x_6); -lean_closure_set(x_22, 6, x_7); -lean_closure_set(x_22, 7, x_15); -lean_closure_set(x_22, 8, x_21); -x_23 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2; +lean_dec(x_16); +x_20 = lean_array_push(x_9, x_19); +x_21 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1___boxed), 15, 9); +lean_closure_set(x_21, 0, x_8); +lean_closure_set(x_21, 1, x_3); +lean_closure_set(x_21, 2, x_2); +lean_closure_set(x_21, 3, x_4); +lean_closure_set(x_21, 4, x_5); +lean_closure_set(x_21, 5, x_6); +lean_closure_set(x_21, 6, x_7); +lean_closure_set(x_21, 7, x_1); +lean_closure_set(x_21, 8, x_20); +x_22 = l_Lean_Meta_withNewEqs_loop___rarg___closed__2; +x_23 = 0; x_24 = 0; -x_25 = 0; -x_26 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_23, x_24, x_19, x_22, x_25, x_10, x_11, x_12, x_13, x_18); -return x_26; +x_25 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_22, x_23, x_18, x_21, x_24, x_10, x_11, x_12, x_13, x_17); +return x_25; } else { -uint8_t x_27; -lean_dec(x_15); +uint8_t x_26; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -2319,34 +2324,35 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_15); +if (x_26 == 0) { -return x_16; +return x_15; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_16, 0); -x_29 = lean_ctor_get(x_16, 1); -lean_inc(x_29); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_15, 0); +x_28 = lean_ctor_get(x_15, 1); lean_inc(x_28); -lean_dec(x_16); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_inc(x_27); +lean_dec(x_15); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; lean_inc(x_6); lean_inc(x_3); -x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__2___boxed), 14, 7); +x_13 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__2), 14, 7); lean_closure_set(x_13, 0, x_1); lean_closure_set(x_13, 1, x_7); lean_closure_set(x_13, 2, x_2); @@ -2354,103 +2360,150 @@ lean_closure_set(x_13, 3, x_3); lean_closure_set(x_13, 4, x_4); lean_closure_set(x_13, 5, x_5); lean_closure_set(x_13, 6, x_6); -x_14 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg(x_6, x_3, x_13, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_Meta_withNewEqs___rarg(x_6, x_3, x_13, x_8, x_9, x_10, x_11, x_12); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; -x_14 = l_Lean_mkAppN(x_1, x_7); -x_15 = l_Lean_LocalDecl_userName(x_2); -x_16 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__3), 12, 6); -lean_closure_set(x_16, 0, x_2); -lean_closure_set(x_16, 1, x_3); -lean_closure_set(x_16, 2, x_7); -lean_closure_set(x_16, 3, x_4); -lean_closure_set(x_16, 4, x_5); -lean_closure_set(x_16, 5, x_6); -x_17 = 0; -x_18 = 0; -x_19 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_15, x_17, x_14, x_16, x_18, x_9, x_10, x_11, x_12, x_13); -return x_19; +lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_14 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__3), 12, 6); +lean_closure_set(x_14, 0, x_1); +lean_closure_set(x_14, 1, x_2); +lean_closure_set(x_14, 2, x_3); +lean_closure_set(x_14, 3, x_4); +lean_closure_set(x_14, 4, x_5); +lean_closure_set(x_14, 5, x_6); +x_15 = 0; +x_16 = 0; +x_17 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__4___rarg(x_8, x_15, x_7, x_14, x_16, x_9, x_10, x_11, x_12, x_13); +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1() { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_14 = lean_array_get_size(x_1); -x_15 = lean_ctor_get(x_2, 2); -x_16 = lean_nat_sub(x_14, x_15); -x_17 = l_Array_extract___rarg(x_1, x_16, x_14); -lean_dec(x_14); -x_18 = lean_ctor_get(x_2, 1); -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Array_extract___rarg(x_1, x_19, x_18); -x_21 = l_Lean_mkAppN(x_3, x_20); -lean_dec(x_20); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Lean_mkAppN(x_1, x_8); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2; +x_17 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_16, x_12, x_13, x_14); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__4(x_2, x_3, x_8, x_4, x_5, x_6, x_15, x_18, x_10, x_11, x_12, x_13, x_19); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_7, 0); +lean_inc(x_21); +lean_dec(x_7); +x_22 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__4(x_2, x_3, x_8, x_4, x_5, x_6, x_15, x_21, x_10, x_11, x_12, x_13, x_14); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_15 = lean_array_get_size(x_1); +x_16 = lean_ctor_get(x_2, 2); +x_17 = lean_nat_sub(x_15, x_16); +x_18 = l_Array_extract___rarg(x_1, x_17, x_15); +lean_dec(x_15); +x_19 = lean_ctor_get(x_2, 1); +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Array_extract___rarg(x_1, x_20, x_19); +x_22 = l_Lean_mkAppN(x_3, x_21); +lean_dec(x_21); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_21); -x_22 = lean_infer_type(x_21, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_22) == 0) +lean_inc(x_22); +x_23 = lean_infer_type(x_22, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__4___boxed), 13, 6); -lean_closure_set(x_25, 0, x_21); -lean_closure_set(x_25, 1, x_4); -lean_closure_set(x_25, 2, x_5); -lean_closure_set(x_25, 3, x_6); -lean_closure_set(x_25, 4, x_7); -lean_closure_set(x_25, 5, x_17); -x_26 = 0; -x_27 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_23, x_25, x_26, x_9, x_10, x_11, x_12, x_24); -return x_27; +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___boxed), 14, 7); +lean_closure_set(x_26, 0, x_22); +lean_closure_set(x_26, 1, x_4); +lean_closure_set(x_26, 2, x_5); +lean_closure_set(x_26, 3, x_6); +lean_closure_set(x_26, 4, x_7); +lean_closure_set(x_26, 5, x_18); +lean_closure_set(x_26, 6, x_8); +x_27 = 0; +x_28 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_24, x_26, x_27, x_10, x_11, x_12, x_13, x_25); +return x_28; } else { -uint8_t x_28; -lean_dec(x_21); -lean_dec(x_17); +uint8_t x_29; +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_28 = !lean_is_exclusive(x_22); -if (x_28 == 0) +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) { -return x_22; +return x_23; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_22, 0); -x_30 = lean_ctor_get(x_22, 1); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_22); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1() { _start: { lean_object* x_1; @@ -2458,89 +2511,90 @@ x_1 = lean_mk_string_unchecked("ill-formed inductive datatype", 29, 29); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_array_get_size(x_1); -x_16 = lean_ctor_get(x_2, 2); -x_17 = lean_ctor_get(x_2, 1); -x_18 = lean_nat_add(x_16, x_17); -x_19 = lean_nat_dec_eq(x_15, x_18); -lean_dec(x_18); -lean_dec(x_15); -if (x_19 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_array_get_size(x_1); +x_17 = lean_ctor_get(x_2, 2); +x_18 = lean_ctor_get(x_2, 1); +x_19 = lean_nat_add(x_17, x_18); +x_20 = lean_nat_dec_eq(x_16, x_19); +lean_dec(x_19); +lean_dec(x_16); +if (x_20 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_20 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4; -x_21 = l_Lean_Meta_throwTacticEx___rarg(x_8, x_5, x_20, x_10, x_11, x_12, x_13, x_14); +x_21 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4; +x_22 = l_Lean_Meta_throwTacticEx___rarg(x_9, x_5, x_21, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -return x_21; +return x_22; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_ctor_get(x_21, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_21); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_dec(x_22); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -lean_object* x_26; lean_object* x_27; -lean_dec(x_8); -x_26 = lean_box(0); -x_27 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_26, x_10, x_11, x_12, x_13, x_14); -return x_27; +lean_object* x_27; lean_object* x_28; +lean_dec(x_9); +x_27 = lean_box(0); +x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27, x_11, x_12, x_13, x_14, x_15); +return x_28; } } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1() { _start: { lean_object* x_1; @@ -2548,36 +2602,36 @@ x_1 = lean_mk_string_unchecked("inductive type expected", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5() { _start: { lean_object* x_1; @@ -2585,189 +2639,194 @@ x_1 = lean_mk_string_unchecked("indexed inductive type expected", 31, 31); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8() { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -switch (lean_obj_tag(x_6)) { +switch (lean_obj_tag(x_7)) { case 4: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec(x_8); -x_14 = lean_ctor_get(x_6, 0); -lean_inc(x_14); -x_15 = lean_st_ref_get(x_12, x_13); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_9); +x_15 = lean_ctor_get(x_7, 0); +lean_inc(x_15); +x_16 = lean_st_ref_get(x_13, x_14); +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 0); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = lean_environment_find(x_18, x_14); -if (lean_obj_tag(x_19) == 0) +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Environment_find_x3f(x_19, x_15); +lean_dec(x_15); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_20; lean_object* x_21; +lean_object* x_21; lean_object* x_22; +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_20 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4; -x_21 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_20, x_9, x_10, x_11, x_12, x_17); +x_21 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4; +x_22 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -return x_21; +return x_22; } else { -lean_object* x_22; -x_22 = lean_ctor_get(x_19, 0); -lean_inc(x_22); -lean_dec(x_19); -if (lean_obj_tag(x_22) == 5) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_23 = lean_ctor_get(x_22, 0); +lean_object* x_23; +x_23 = lean_ctor_get(x_20, 0); lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_ctor_get(x_23, 2); -lean_inc(x_24); -x_25 = lean_unsigned_to_nat(0u); -x_26 = lean_nat_dec_lt(x_25, x_24); -lean_dec(x_24); -if (x_26 == 0) +lean_dec(x_20); +if (lean_obj_tag(x_23) == 5) { -lean_object* x_27; lean_object* x_28; uint8_t x_29; +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); lean_dec(x_23); +x_25 = lean_ctor_get(x_24, 2); +lean_inc(x_25); +x_26 = lean_unsigned_to_nat(0u); +x_27 = lean_nat_dec_lt(x_26, x_25); +lean_dec(x_25); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_dec(x_24); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_27 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8; -x_28 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_27, x_9, x_10, x_11, x_12, x_17); +x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8; +x_29 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_28, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) { -return x_28; +return x_29; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ctor_get(x_28, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_28); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_dec(x_29); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } else { -lean_object* x_33; lean_object* x_34; -x_33 = lean_box(0); -x_34 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6(x_7, x_23, x_6, x_5, x_1, x_2, x_3, x_4, x_33, x_9, x_10, x_11, x_12, x_17); -lean_dec(x_23); -lean_dec(x_7); -return x_34; +lean_object* x_34; lean_object* x_35; +x_34 = lean_box(0); +x_35 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7(x_8, x_24, x_7, x_2, x_1, x_4, x_5, x_3, x_6, x_34, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_24); +lean_dec(x_8); +return x_35; } } else { -lean_object* x_35; lean_object* x_36; -lean_dec(x_22); +lean_object* x_36; lean_object* x_37; +lean_dec(x_23); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_35 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4; -x_36 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_35, x_9, x_10, x_11, x_12, x_17); +x_36 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4; +x_37 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_36, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -return x_36; +return x_37; } } } case 5: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = lean_ctor_get(x_6, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_6, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = lean_ctor_get(x_7, 0); lean_inc(x_38); -lean_dec(x_6); -x_39 = lean_array_set(x_7, x_8, x_38); -x_40 = lean_unsigned_to_nat(1u); -x_41 = lean_nat_sub(x_8, x_40); -lean_dec(x_8); -x_6 = x_37; -x_7 = x_39; -x_8 = x_41; +x_39 = lean_ctor_get(x_7, 1); +lean_inc(x_39); +lean_dec(x_7); +x_40 = lean_array_set(x_8, x_9, x_39); +x_41 = lean_unsigned_to_nat(1u); +x_42 = lean_nat_sub(x_9, x_41); +lean_dec(x_9); +x_7 = x_38; +x_8 = x_40; +x_9 = x_42; goto _start; } default: { -lean_object* x_43; lean_object* x_44; +lean_object* x_44; lean_object* x_45; +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_43 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4; -x_44 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_43, x_9, x_10, x_11, x_12, x_13); +x_44 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4; +x_45 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_44, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -return x_44; +return x_45; } } } } -static lean_object* _init_l_Lean_Meta_generalizeIndices___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -2775,53 +2834,56 @@ x_1 = lean_mk_string_unchecked("generalizeIndices", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Meta_generalizeIndices___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_generalizeIndices___lambda__1___closed__1; +x_2 = l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices_x27___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_8 = lean_ctor_get(x_3, 2); -lean_inc(x_8); -x_9 = l_Lean_Meta_getLocalInstances(x_3, x_4, x_5, x_6, x_7); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_4, 2); +lean_inc(x_9); +x_10 = l_Lean_Meta_getLocalInstances(x_4, x_5, x_6, x_7, x_8); +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Meta_generalizeIndices___lambda__1___closed__2; +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2; lean_inc(x_1); -x_13 = l_Lean_MVarId_checkNotAssigned(x_1, x_12, x_3, x_4, x_5, x_6, x_11); -if (lean_obj_tag(x_13) == 0) +x_14 = l_Lean_MVarId_checkNotAssigned(x_1, x_13, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -lean_inc(x_3); -x_15 = l_Lean_FVarId_getDecl(x_2, x_3, x_4, x_5, x_6, x_14); -if (lean_obj_tag(x_15) == 0) +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +x_16 = lean_infer_type(x_2, x_4, x_5, x_6, x_7, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_LocalDecl_type(x_16); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_19 = lean_whnf(x_18, x_3, x_4, x_5, x_6, x_17); +x_19 = l_Lean_Meta_whnfD(x_17, x_4, x_5, x_6, x_7, x_18); if (lean_obj_tag(x_19) == 0) { lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; @@ -2838,19 +2900,20 @@ x_25 = lean_mk_array(x_23, x_24); x_26 = lean_unsigned_to_nat(1u); x_27 = lean_nat_sub(x_23, x_26); lean_dec(x_23); -x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1(x_1, x_8, x_10, x_12, x_16, x_20, x_25, x_27, x_3, x_4, x_5, x_6, x_21); +x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1(x_1, x_2, x_3, x_9, x_11, x_13, x_20, x_25, x_27, x_4, x_5, x_6, x_7, x_21); return x_28; } else { uint8_t x_29; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_8); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); x_29 = !lean_is_exclusive(x_19); if (x_29 == 0) @@ -2875,26 +2938,28 @@ return x_32; else { uint8_t x_33; -lean_dec(x_10); -lean_dec(x_8); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_15); +x_33 = !lean_is_exclusive(x_16); if (x_33 == 0) { -return x_15; +return x_16; } else { lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_15, 0); -x_35 = lean_ctor_get(x_15, 1); +x_34 = lean_ctor_get(x_16, 0); +x_35 = lean_ctor_get(x_16, 1); lean_inc(x_35); lean_inc(x_34); -lean_dec(x_15); +lean_dec(x_16); x_36 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_36, 0, x_34); lean_ctor_set(x_36, 1, x_35); @@ -2905,27 +2970,28 @@ return x_36; else { uint8_t x_37; -lean_dec(x_10); -lean_dec(x_8); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_37 = !lean_is_exclusive(x_13); +x_37 = !lean_is_exclusive(x_14); if (x_37 == 0) { -return x_13; +return x_14; } else { lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_13, 0); -x_39 = lean_ctor_get(x_13, 1); +x_38 = lean_ctor_get(x_14, 0); +x_39 = lean_ctor_get(x_14, 1); lean_inc(x_39); lean_inc(x_38); -lean_dec(x_13); +lean_dec(x_14); x_40 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); @@ -2934,67 +3000,122 @@ return x_40; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; lean_object* x_9; +lean_object* x_9; lean_object* x_10; lean_inc(x_1); -x_8 = lean_alloc_closure((void*)(l_Lean_Meta_generalizeIndices___lambda__1), 7, 2); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_2); -x_9 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -return x_9; +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_generalizeIndices_x27___lambda__1), 8, 3); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_2); +lean_closure_set(x_9, 2, x_3); +x_10 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_9); lean_dec(x_7); lean_dec(x_4); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { lean_object* x_15; -x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_1); +x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_9); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; -x_14 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); -return x_14; +lean_object* x_15; +x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_14; -x_14 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); +lean_object* x_16; +x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_10); lean_dec(x_2); lean_dec(x_1); -return x_14; +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_15; -x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_8; +lean_inc(x_3); +x_8 = l_Lean_FVarId_getDecl(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_LocalDecl_toExpr(x_9); +x_12 = l_Lean_LocalDecl_userName(x_9); lean_dec(x_9); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = l_Lean_Meta_generalizeIndices_x27(x_2, x_11, x_13, x_3, x_4, x_5, x_6, x_10); +return x_14; +} +else +{ +uint8_t x_15; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_15; +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_8); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_generalizeIndices(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +lean_inc(x_1); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_generalizeIndices___lambda__1), 7, 2); +lean_closure_set(x_8, 0, x_2); +lean_closure_set(x_8, 1, x_1); +x_9 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; } } static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1() { @@ -3024,7 +3145,8 @@ x_14 = lean_ctor_get(x_12, 0); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); lean_dec(x_14); -x_16 = lean_environment_find(x_15, x_11); +x_16 = l_Lean_Environment_find_x3f(x_15, x_11); +lean_dec(x_11); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; @@ -3081,7 +3203,8 @@ lean_inc(x_27); lean_dec(x_26); x_28 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1; x_29 = l_Lean_Name_str___override(x_27, x_28); -x_30 = lean_environment_find(x_1, x_29); +x_30 = l_Lean_Environment_find_x3f(x_1, x_29); +lean_dec(x_29); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; @@ -3223,7 +3346,8 @@ lean_dec(x_12); x_55 = lean_ctor_get(x_53, 0); lean_inc(x_55); lean_dec(x_53); -x_56 = lean_environment_find(x_55, x_11); +x_56 = l_Lean_Environment_find_x3f(x_55, x_11); +lean_dec(x_11); if (lean_obj_tag(x_56) == 0) { lean_object* x_57; lean_object* x_58; @@ -3284,7 +3408,8 @@ lean_inc(x_69); lean_dec(x_68); x_70 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1; x_71 = l_Lean_Name_str___override(x_69, x_70); -x_72 = lean_environment_find(x_1, x_71); +x_72 = l_Lean_Environment_find_x3f(x_1, x_71); +lean_dec(x_71); if (lean_obj_tag(x_72) == 0) { lean_object* x_73; lean_object* x_74; @@ -13520,7 +13645,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_11 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_11 = l_Lean_Meta_withNewEqs___rarg___closed__1; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_7); @@ -13539,7 +13664,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_14 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_14 = l_Lean_Meta_withNewEqs___rarg___closed__1; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_7); @@ -13551,7 +13676,7 @@ size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; x_16 = 0; x_17 = lean_usize_of_nat(x_8); lean_dec(x_8); -x_18 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_18 = l_Lean_Meta_withNewEqs___rarg___closed__1; x_19 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_unifyCasesEqs___spec__2(x_1, x_2, x_16, x_17, x_18, x_3, x_4, x_5, x_6, x_7); return x_19; } @@ -15313,7 +15438,7 @@ lean_dec(x_35); x_41 = l_Lean_LocalDecl_fvarId(x_34); lean_dec(x_34); x_42 = lean_box(0); -x_43 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_43 = l_Lean_Meta_withNewEqs___rarg___closed__1; lean_inc(x_2); x_44 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1), 9, 4); lean_closure_set(x_44, 0, x_2); @@ -15497,7 +15622,7 @@ lean_dec(x_66); x_72 = l_Lean_LocalDecl_fvarId(x_65); lean_dec(x_65); x_73 = lean_box(0); -x_74 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_74 = l_Lean_Meta_withNewEqs___rarg___closed__1; lean_inc(x_2); x_75 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1), 9, 4); lean_closure_set(x_75, 0, x_2); @@ -15999,7 +16124,7 @@ lean_dec(x_39); x_45 = l_Lean_LocalDecl_fvarId(x_38); lean_dec(x_38); x_46 = lean_box(0); -x_47 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_47 = l_Lean_Meta_withNewEqs___rarg___closed__1; lean_inc(x_2); x_48 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1), 9, 4); lean_closure_set(x_48, 0, x_2); @@ -16183,7 +16308,7 @@ lean_dec(x_70); x_76 = l_Lean_LocalDecl_fvarId(x_69); lean_dec(x_69); x_77 = lean_box(0); -x_78 = l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1; +x_78 = l_Lean_Meta_withNewEqs___rarg___closed__1; lean_inc(x_2); x_79 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_casesRec___spec__6___lambda__1), 9, 4); lean_closure_set(x_79, 0, x_2); @@ -18325,7 +18450,7 @@ return x_109; } } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1() { _start: { lean_object* x_1; @@ -18333,27 +18458,27 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2; x_2 = l_Lean_Meta_Cases_cases___lambda__2___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4() { _start: { lean_object* x_1; @@ -18361,17 +18486,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6() { _start: { lean_object* x_1; @@ -18379,47 +18504,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8; x_2 = l_Lean_Meta_Cases_cases___lambda__2___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9; x_2 = l_Lean_Meta_Cases_cases___lambda__2___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11() { _start: { lean_object* x_1; @@ -18427,17 +18552,17 @@ x_1 = lean_mk_string_unchecked("Cases", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13() { _start: { lean_object* x_1; @@ -18445,33 +18570,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14; -x_2 = lean_unsigned_to_nat(3794u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14; +x_2 = lean_unsigned_to_nat(3900u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_Cases_cases___lambda__2___closed__8; x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -18533,12 +18658,12 @@ l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__6 = _init lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__6); l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__7 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__7(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__7); -l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__1); -l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs_loop___rarg___closed__2); -l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1 = _init_l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_withNewEqs___rarg___closed__1); +l_Lean_Meta_withNewEqs_loop___rarg___closed__1 = _init_l_Lean_Meta_withNewEqs_loop___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_withNewEqs_loop___rarg___closed__1); +l_Lean_Meta_withNewEqs_loop___rarg___closed__2 = _init_l_Lean_Meta_withNewEqs_loop___rarg___closed__2(); +lean_mark_persistent(l_Lean_Meta_withNewEqs_loop___rarg___closed__2); +l_Lean_Meta_withNewEqs___rarg___closed__1 = _init_l_Lean_Meta_withNewEqs___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_withNewEqs___rarg___closed__1); l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__1 = _init_l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__1); l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__2 = _init_l_Lean_Meta_generalizeTargetsEq___lambda__3___closed__2(); @@ -18551,34 +18676,38 @@ l_Lean_Meta_generalizeTargetsEq___closed__1 = _init_l_Lean_Meta_generalizeTarget lean_mark_persistent(l_Lean_Meta_generalizeTargetsEq___closed__1); l_Lean_Meta_generalizeTargetsEq___closed__2 = _init_l_Lean_Meta_generalizeTargetsEq___closed__2(); lean_mark_persistent(l_Lean_Meta_generalizeTargetsEq___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__1); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__3); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___lambda__6___closed__4); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__3); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__4); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__5); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__6); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__7); -l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__8); -l_Lean_Meta_generalizeIndices___lambda__1___closed__1 = _init_l_Lean_Meta_generalizeIndices___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_generalizeIndices___lambda__1___closed__1); -l_Lean_Meta_generalizeIndices___lambda__1___closed__2 = _init_l_Lean_Meta_generalizeIndices___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_generalizeIndices___lambda__1___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__1); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__5___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__1); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__3); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___lambda__7___closed__4); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__1); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__2); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__3); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__4); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__5); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__6); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__7); +l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8(); +lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices_x27___spec__1___closed__8); +l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1 = _init_l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__1); +l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2 = _init_l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_generalizeIndices_x27___lambda__1___closed__2); l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1 = _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1(); lean_mark_persistent(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_mkCasesContext_x3f___spec__1___closed__1); l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__1 = _init_l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__56___closed__1(); @@ -18669,37 +18798,37 @@ l_Lean_MVarId_byCasesDec___closed__4 = _init_l_Lean_MVarId_byCasesDec___closed__ lean_mark_persistent(l_Lean_MVarId_byCasesDec___closed__4); l_Lean_MVarId_byCasesDec___closed__5 = _init_l_Lean_MVarId_byCasesDec___closed__5(); lean_mark_persistent(l_Lean_MVarId_byCasesDec___closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__13); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__14); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794____closed__15); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3794_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__13); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__14); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900____closed__15); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Cases___hyg_3900_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Constructor.c b/stage0/stdlib/Lean/Meta/Tactic/Constructor.c index 521ffe434ecb..4415f2f1b40b 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Constructor.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Constructor.c @@ -27,7 +27,7 @@ lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*) lean_object* lean_mk_array(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_constructor___lambda__2___closed__3; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_MVarId_existsIntro___lambda__2___closed__7; static lean_object* l_Lean_MVarId_constructor___lambda__2___closed__4; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_constructor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -487,7 +487,8 @@ lean_dec(x_17); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_find(x_20, x_15); +x_21 = l_Lean_Environment_find_x3f(x_20, x_15); +lean_dec(x_15); if (lean_obj_tag(x_21) == 0) { lean_object* x_22; lean_object* x_23; @@ -1233,7 +1234,8 @@ lean_dec(x_17); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_find(x_20, x_15); +x_21 = l_Lean_Environment_find_x3f(x_20, x_15); +lean_dec(x_15); if (lean_obj_tag(x_21) == 0) { lean_object* x_22; lean_object* x_23; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c b/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c index 9ee1481ae4ee..297f72d9056f 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c @@ -71,7 +71,7 @@ lean_object* lean_array_fget(lean_object*, lean_object*); uint8_t lean_expr_has_loose_bvar(lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_commitWhen___at_Lean_Meta_ElimEmptyInductive_elim___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Meta_ElimEmptyInductive_elim___spec__6___closed__3; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Contradiction___hyg_4437____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_ElimEmptyInductive_elim___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -732,7 +732,8 @@ x_18 = lean_ctor_get(x_16, 0); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); lean_dec(x_18); -x_20 = lean_environment_find(x_19, x_15); +x_20 = l_Lean_Environment_find_x3f(x_19, x_15); +lean_dec(x_15); if (lean_obj_tag(x_20) == 0) { uint8_t x_21; lean_object* x_22; @@ -804,7 +805,8 @@ lean_dec(x_16); x_38 = lean_ctor_get(x_36, 0); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_environment_find(x_38, x_15); +x_39 = l_Lean_Environment_find_x3f(x_38, x_15); +lean_dec(x_15); if (lean_obj_tag(x_39) == 0) { uint8_t x_40; lean_object* x_41; lean_object* x_42; @@ -918,7 +920,8 @@ if (lean_is_exclusive(x_65)) { x_69 = lean_ctor_get(x_66, 0); lean_inc(x_69); lean_dec(x_66); -x_70 = lean_environment_find(x_69, x_64); +x_70 = l_Lean_Environment_find_x3f(x_69, x_64); +lean_dec(x_64); if (lean_obj_tag(x_70) == 0) { uint8_t x_71; lean_object* x_72; lean_object* x_73; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Delta.c b/stage0/stdlib/Lean/Meta/Tactic/Delta.c index b4e1ff9cc76f..de8bdb7ba630 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Delta.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Delta.c @@ -21,7 +21,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_deltaExpand___lambda__1___boxed(lean_object static lean_object* l_Lean_Meta_deltaExpand___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_deltaTarget___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_name(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t l_Lean_ConstantInfo_hasValue(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_deltaExpand___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_deltaTarget___closed__2; @@ -71,7 +71,8 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -x_14 = lean_environment_find(x_13, x_7); +x_14 = l_Lean_Environment_find_x3f(x_13, x_7); +lean_dec(x_7); if (lean_obj_tag(x_14) == 0) { lean_object* x_15; @@ -355,7 +356,8 @@ lean_dec(x_9); x_82 = lean_ctor_get(x_80, 0); lean_inc(x_82); lean_dec(x_80); -x_83 = lean_environment_find(x_82, x_7); +x_83 = l_Lean_Environment_find_x3f(x_82, x_7); +lean_dec(x_7); if (lean_obj_tag(x_83) == 0) { lean_object* x_84; lean_object* x_85; diff --git a/stage0/stdlib/Lean/Meta/Tactic/FunInd.c b/stage0/stdlib/Lean/Meta/Tactic/FunInd.c index 9ff57d2014b6..051b22c95c2a 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/FunInd.c +++ b/stage0/stdlib/Lean/Meta/Tactic/FunInd.c @@ -241,7 +241,7 @@ static lean_object* l_Lean_Tactic_FunInd_abstractIndependentMVars___closed__3; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at_Lean_Tactic_FunInd_unpackMutualInduction___spec__1___lambda__3___closed__2; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_withLetDecls___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2196,8 +2196,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -2241,8 +2240,7 @@ lean_dec(x_8); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -lean_inc(x_1); -x_26 = lean_environment_find(x_25, x_1); +x_26 = l_Lean_Environment_find_x3f(x_25, x_1); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; @@ -28568,8 +28566,7 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -lean_inc(x_1); -x_14 = lean_environment_find(x_13, x_1); +x_14 = l_Lean_Environment_find_x3f(x_13, x_1); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -28617,8 +28614,7 @@ lean_dec(x_9); x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); -lean_inc(x_1); -x_28 = lean_environment_find(x_27, x_1); +x_28 = l_Lean_Environment_find_x3f(x_27, x_1); if (lean_obj_tag(x_28) == 0) { uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind.c b/stage0/stdlib/Lean/Meta/Tactic/Grind.c index 08ec47f8a228..6ba86e8a4b51 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind -// Imports: Lean.Meta.Tactic.Grind.Attr Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Canon Lean.Meta.Tactic.Grind.MarkNestedProofs Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.Propagate Lean.Meta.Tactic.Grind.PP Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Ctor Lean.Meta.Tactic.Grind.Parser Lean.Meta.Tactic.Grind.EMatchTheorem Lean.Meta.Tactic.Grind.EMatch Lean.Meta.Tactic.Grind.Main +// Imports: Lean.Meta.Tactic.Grind.Attr Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Canon Lean.Meta.Tactic.Grind.MarkNestedProofs Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.Propagate Lean.Meta.Tactic.Grind.PP Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Ctor Lean.Meta.Tactic.Grind.Parser Lean.Meta.Tactic.Grind.EMatchTheorem Lean.Meta.Tactic.Grind.EMatch Lean.Meta.Tactic.Grind.Main Lean.Meta.Tactic.Grind.CasesMatch #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,86 +13,113 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__8; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__9; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597_(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__14; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__10; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__3; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__13; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__11; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__1; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152_(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__6; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78_(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__5; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__12; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__2; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__16; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_41_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__15; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_(lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__18; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4_(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__7; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__4; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_115____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485_(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__2; static lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_78____closed__1; @@ -473,19 +500,20 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("instance", 8, 8); +x_1 = lean_mk_string_unchecked("search", 6, 6); return x_1; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_189____closed__1; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__3() { @@ -513,20 +541,19 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("assignment", 10, 10); +x_1 = lean_mk_string_unchecked("instance", 8, 8); return x_1; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_226____closed__1; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__3() { @@ -554,18 +581,20 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("issues", 6, 6); +x_1 = lean_mk_string_unchecked("assignment", 10, 10); return x_1; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_152____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_263____closed__1; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_300____closed__3() { @@ -593,7 +622,7 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("simp", 4, 4); +x_1 = lean_mk_string_unchecked("issues", 6, 6); return x_1; } } @@ -628,161 +657,160 @@ x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("debug", 5, 5); +x_1 = lean_mk_string_unchecked("simp", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; -x_2 = lean_unsigned_to_nat(375u); +x_2 = lean_unsigned_to_nat(374u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("proofs", 6, 6); +x_1 = lean_mk_string_unchecked("split", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; -x_2 = lean_unsigned_to_nat(412u); +x_2 = lean_unsigned_to_nat(411u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("congr", 5, 5); +x_1 = lean_mk_string_unchecked("candidate", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; -x_2 = lean_unsigned_to_nat(449u); +x_2 = lean_unsigned_to_nat(448u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("proof", 5, 5); +x_1 = lean_mk_string_unchecked("resolved", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; -x_2 = lean_unsigned_to_nat(486u); +x_2 = lean_unsigned_to_nat(485u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -791,19 +819,18 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("proj", 4, 4); +x_1 = lean_mk_string_unchecked("debug", 5, 5); return x_1; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; -x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__3() { @@ -831,7 +858,7 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("parent", 6, 6); +x_1 = lean_mk_string_unchecked("proofs", 6, 6); return x_1; } } @@ -840,7 +867,7 @@ static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -867,6 +894,278 @@ x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("congr", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(597u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proof", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(634u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proj", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(671u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("parent", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(708u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("final", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(745u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("forallPropagator", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(782u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__17; +x_2 = lean_unsigned_to_nat(819u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1; +x_3 = 0; +x_4 = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} lean_object* initialize_Lean_Meta_Tactic_Grind_Attr(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_RevertAll(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); @@ -886,6 +1185,7 @@ lean_object* initialize_Lean_Meta_Tactic_Grind_Parser(uint8_t builtin, lean_obje lean_object* initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Main(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_CasesMatch(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind(uint8_t builtin, lean_object* w) { lean_object * res; @@ -948,6 +1248,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Main(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_CasesMatch(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1(); lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__1); l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_4____closed__2(); @@ -1068,40 +1371,40 @@ lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337____cl if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_337_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_375_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_374_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_412_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_411_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_449_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_448_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__1); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__2); -l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_486_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_485_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_523____closed__1(); @@ -1122,6 +1425,67 @@ lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560____cl if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_560_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_597_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_634_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_671_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_708_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_745_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__2); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_782_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__1); +l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819____closed__2); +if (builtin) {res = l_Lean_initFn____x40_Lean_Meta_Tactic_Grind___hyg_819_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c index 29d5704f10e6..6e7594c01188 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Canon.c @@ -14,10 +14,11 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__16___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_canonElemCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -33,6 +34,7 @@ lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, l size_t lean_uint64_to_usize(uint64_t); uint64_t lean_uint64_lor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -41,36 +43,37 @@ lean_object* lean_array_push(lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___lambda__1___boxed(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__11(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__16(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__1; +static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkPtrMap___rarg(lean_object*); -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -79,7 +82,7 @@ uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__8; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___closed__2; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3; lean_object* l_Lean_Meta_getFunInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_usize_to_uint64(size_t); lean_object* lean_nat_div(lean_object*, lean_object*); @@ -88,11 +91,17 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_Canon_instInhabitedShouldCanonResult; @@ -103,17 +112,16 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2(lean_ static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__7___closed__1; uint64_t l_Lean_Expr_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_toCtorIdx(uint8_t); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_ParamInfo_isInstImplicit(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___boxed(lean_object**); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); uint8_t l_Lean_Expr_fvarsSubset(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_shouldCanon___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -132,7 +140,7 @@ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_Canon_c lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3; +static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3; static lean_object* l_Lean_Meta_Grind_Canon_canonElemCore___lambda__2___closed__1; static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___closed__1; static lean_object* l_Lean_Meta_Grind_Canon_instInhabitedState___closed__2; @@ -144,9 +152,7 @@ lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMe uint64_t lean_uint64_xor(uint64_t, uint64_t); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_Canon_canonElemCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1; lean_object* lean_nat_mul(lean_object*, lean_object*); lean_object* l_Lean_Meta_isTypeFormer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -167,9 +173,8 @@ size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_instInhabitedState; @@ -183,21 +188,25 @@ static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCo static uint64_t l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); static lean_object* l_Lean_Meta_Grind_Canon_instInhabitedState___closed__3; +static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__3; lean_object* l_Lean_isDiagnosticsEnabled(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isForall(lean_object*); static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2; static lean_object* l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_toCtorIdx___boxed(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; lean_object* l_ReaderT_instMonad___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_Canon_canonElemCore___spec__2___closed__2; +uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(uint8_t, uint8_t); static lean_object* _init_l_Lean_Meta_Grind_Canon_instInhabitedState___closed__1() { _start: { @@ -5247,45 +5256,7 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_StateT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2; -x_2 = l_Lean_instInhabitedExpr; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3; -x_10 = lean_panic_fn(x_9, x_1); -x_11 = lean_apply_7(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; -} -} -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -5316,7 +5287,7 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -5399,7 +5370,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -5479,7 +5450,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; @@ -5498,7 +5469,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = lean_array_fget(x_2, x_1); x_7 = lean_box(0); x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_3, x_6); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__5(x_3, x_6); x_10 = lean_unsigned_to_nat(1u); x_11 = lean_nat_add(x_1, x_10); lean_dec(x_1); @@ -5509,7 +5480,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; @@ -5520,11 +5491,11 @@ lean_dec(x_2); x_5 = lean_box(0); x_6 = lean_mk_array(x_4, x_5); x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__4(x_7, x_1, x_6); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(x_7, x_1, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -5551,7 +5522,7 @@ x_11 = lean_usize_dec_eq(x_9, x_10); if (x_11 == 0) { lean_object* x_12; -x_12 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_8); +x_12 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_8); lean_ctor_set(x_3, 2, x_12); return x_3; } @@ -5580,7 +5551,7 @@ x_18 = lean_usize_dec_eq(x_16, x_17); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; -x_19 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_15); +x_19 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_15); x_20 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_20, 0, x_13); lean_ctor_set(x_20, 1, x_14); @@ -5602,6 +5573,44 @@ return x_21; } } } +static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_StateT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3; +x_10 = lean_panic_fn(x_9, x_1); +x_11 = lean_apply_7(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -5612,7 +5621,7 @@ x_15 = lean_usize_dec_eq(x_13, x_14); if (x_15 == 0) { lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_array_set(x_3, x_2, x_5); +x_16 = lean_array_fset(x_3, x_2, x_5); x_17 = 1; x_18 = lean_box(x_17); x_19 = lean_alloc_ctor(0, 2, 0); @@ -5648,648 +5657,606 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_4, 1); -x_17 = lean_nat_dec_lt(x_6, x_16); -if (x_17 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_6, 1); +x_19 = lean_nat_dec_lt(x_8, x_18); +if (x_19 == 0) { -lean_object* x_18; lean_object* x_19; +lean_object* x_20; lean_object* x_21; +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_6); +lean_dec(x_8); lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_5); -lean_ctor_set(x_18, 1, x_10); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_15); -return x_19; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_202; uint8_t x_203; -x_29 = lean_ctor_get(x_5, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_5, 1); -lean_inc(x_30); -lean_dec(x_5); -x_202 = lean_array_get_size(x_29); -x_203 = lean_nat_dec_lt(x_6, x_202); -lean_dec(x_202); -if (x_203 == 0) -{ -lean_object* x_204; lean_object* x_205; -x_204 = l_Lean_instInhabitedExpr; -x_205 = l_outOfBounds___rarg(x_204); -x_31 = x_205; -goto block_201; +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_7); +lean_ctor_set(x_20, 1, x_12); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_17); +return x_21; } else { -lean_object* x_206; -x_206 = lean_array_fget(x_29, x_6); -x_31 = x_206; -goto block_201; -} -block_28: -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_ctor_get(x_22, 0); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_4, 2); -x_26 = lean_nat_add(x_6, x_25); -lean_dec(x_6); -x_5 = x_24; -x_6 = x_26; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_10 = x_23; -x_15 = x_21; -goto _start; -} -block_201: -{ -lean_object* x_32; +lean_object* x_22; lean_object* x_23; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_7, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_7, 1); +lean_inc(x_32); +lean_dec(x_7); +x_33 = lean_array_fget(x_31, x_8); +lean_inc(x_16); +lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_31); -x_32 = l_Lean_Meta_Grind_Canon_shouldCanon(x_2, x_6, x_31, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); -x_34 = lean_unbox(x_33); -lean_dec(x_33); -switch (x_34) { -case 0: +x_34 = l_Lean_Meta_Grind_Canon_shouldCanon(x_3, x_8, x_33, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_35; lean_object* x_36; uint64_t x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_35 = lean_ctor_get(x_11, 0); +lean_object* x_35; uint8_t x_36; +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -x_36 = lean_ctor_get(x_32, 1); -lean_inc(x_36); -lean_dec(x_32); -x_37 = lean_ctor_get_uint64(x_11, sizeof(void*)*7); -x_38 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 8); -x_39 = lean_ctor_get(x_11, 1); -lean_inc(x_39); -x_40 = lean_ctor_get(x_11, 2); -lean_inc(x_40); -x_41 = lean_ctor_get(x_11, 3); +x_36 = lean_unbox(x_35); +lean_dec(x_35); +switch (x_36) { +case 0: +{ +lean_object* x_37; lean_object* x_38; uint64_t x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_37 = lean_ctor_get(x_13, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_34, 1); +lean_inc(x_38); +lean_dec(x_34); +x_39 = lean_ctor_get_uint64(x_13, sizeof(void*)*7); +x_40 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 8); +x_41 = lean_ctor_get(x_13, 1); lean_inc(x_41); -x_42 = lean_ctor_get(x_11, 4); +x_42 = lean_ctor_get(x_13, 2); lean_inc(x_42); -x_43 = lean_ctor_get(x_11, 5); +x_43 = lean_ctor_get(x_13, 3); lean_inc(x_43); -x_44 = lean_ctor_get(x_11, 6); +x_44 = lean_ctor_get(x_13, 4); lean_inc(x_44); -x_45 = !lean_is_exclusive(x_35); -if (x_45 == 0) -{ -uint8_t x_46; uint8_t x_47; uint8_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; -x_46 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 9); -x_47 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 10); -x_48 = 1; -lean_ctor_set_uint8(x_35, 9, x_48); -x_49 = 2; -x_50 = lean_uint64_shift_right(x_37, x_49); -x_51 = lean_uint64_shift_left(x_50, x_49); -x_52 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; -x_53 = lean_uint64_lor(x_51, x_52); -x_54 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_54, 0, x_35); -lean_ctor_set(x_54, 1, x_39); -lean_ctor_set(x_54, 2, x_40); -lean_ctor_set(x_54, 3, x_41); -lean_ctor_set(x_54, 4, x_42); -lean_ctor_set(x_54, 5, x_43); -lean_ctor_set(x_54, 6, x_44); -lean_ctor_set_uint64(x_54, sizeof(void*)*7, x_53); -lean_ctor_set_uint8(x_54, sizeof(void*)*7 + 8, x_38); -lean_ctor_set_uint8(x_54, sizeof(void*)*7 + 9, x_46); -lean_ctor_set_uint8(x_54, sizeof(void*)*7 + 10, x_47); -x_55 = 0; +x_45 = lean_ctor_get(x_13, 5); +lean_inc(x_45); +x_46 = lean_ctor_get(x_13, 6); +lean_inc(x_46); +x_47 = !lean_is_exclusive(x_37); +if (x_47 == 0) +{ +uint8_t x_48; uint8_t x_49; uint8_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; uint64_t x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; +x_48 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 9); +x_49 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 10); +x_50 = 1; +lean_ctor_set_uint8(x_37, 9, x_50); +x_51 = 2; +x_52 = lean_uint64_shift_right(x_39, x_51); +x_53 = lean_uint64_shift_left(x_52, x_51); +x_54 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; +x_55 = lean_uint64_lor(x_53, x_54); +x_56 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_56, 0, x_37); +lean_ctor_set(x_56, 1, x_41); +lean_ctor_set(x_56, 2, x_42); +lean_ctor_set(x_56, 3, x_43); +lean_ctor_set(x_56, 4, x_44); +lean_ctor_set(x_56, 5, x_45); +lean_ctor_set(x_56, 6, x_46); +lean_ctor_set_uint64(x_56, sizeof(void*)*7, x_55); +lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 8, x_40); +lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 9, x_48); +lean_ctor_set_uint8(x_56, sizeof(void*)*7 + 10, x_49); +x_57 = 0; +lean_inc(x_16); +lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_31); -lean_inc(x_6); +lean_inc(x_33); +lean_inc(x_8); lean_inc(x_1); -x_56 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_6, x_31, x_55, x_10, x_54, x_12, x_13, x_14, x_36); -if (lean_obj_tag(x_56) == 0) +x_58 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_8, x_33, x_57, x_12, x_56, x_14, x_15, x_16, x_38); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = lean_ctor_get(x_57, 0); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); -x_60 = lean_ctor_get(x_57, 1); +x_60 = lean_ctor_get(x_58, 1); lean_inc(x_60); -lean_dec(x_57); -x_61 = lean_unbox(x_30); -lean_dec(x_30); -x_62 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_61, x_59, x_9, x_60, x_11, x_12, x_13, x_14, x_58); -lean_dec(x_31); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_20 = x_63; -x_21 = x_64; -goto block_28; +lean_dec(x_58); +x_61 = lean_ctor_get(x_59, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +lean_dec(x_59); +x_63 = lean_unbox(x_32); +lean_dec(x_32); +x_64 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_33, x_8, x_31, x_63, x_61, x_11, x_62, x_13, x_14, x_15, x_16, x_60); +lean_dec(x_33); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_22 = x_65; +x_23 = x_66; +goto block_30; } else { -uint8_t x_65; +uint8_t x_67; +lean_dec(x_33); +lean_dec(x_32); lean_dec(x_31); -lean_dec(x_30); -lean_dec(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_9); -lean_dec(x_6); +lean_dec(x_8); lean_dec(x_1); -x_65 = !lean_is_exclusive(x_56); -if (x_65 == 0) +x_67 = !lean_is_exclusive(x_58); +if (x_67 == 0) { -return x_56; +return x_58; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_56, 0); -x_67 = lean_ctor_get(x_56, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_56); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -return x_68; -} -} -} -else -{ -uint8_t x_69; uint8_t x_70; uint8_t x_71; uint8_t x_72; uint8_t x_73; uint8_t x_74; uint8_t x_75; uint8_t x_76; uint8_t x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; uint8_t x_82; uint8_t x_83; uint8_t x_84; uint8_t x_85; uint8_t x_86; uint8_t x_87; lean_object* x_88; uint64_t x_89; uint64_t x_90; uint64_t x_91; uint64_t x_92; uint64_t x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; -x_69 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 9); -x_70 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 10); -x_71 = lean_ctor_get_uint8(x_35, 0); -x_72 = lean_ctor_get_uint8(x_35, 1); -x_73 = lean_ctor_get_uint8(x_35, 2); -x_74 = lean_ctor_get_uint8(x_35, 3); -x_75 = lean_ctor_get_uint8(x_35, 4); -x_76 = lean_ctor_get_uint8(x_35, 5); -x_77 = lean_ctor_get_uint8(x_35, 6); -x_78 = lean_ctor_get_uint8(x_35, 7); -x_79 = lean_ctor_get_uint8(x_35, 8); -x_80 = lean_ctor_get_uint8(x_35, 10); -x_81 = lean_ctor_get_uint8(x_35, 11); -x_82 = lean_ctor_get_uint8(x_35, 12); -x_83 = lean_ctor_get_uint8(x_35, 13); -x_84 = lean_ctor_get_uint8(x_35, 14); -x_85 = lean_ctor_get_uint8(x_35, 15); -x_86 = lean_ctor_get_uint8(x_35, 16); -lean_dec(x_35); -x_87 = 1; -x_88 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_88, 0, x_71); -lean_ctor_set_uint8(x_88, 1, x_72); -lean_ctor_set_uint8(x_88, 2, x_73); -lean_ctor_set_uint8(x_88, 3, x_74); -lean_ctor_set_uint8(x_88, 4, x_75); -lean_ctor_set_uint8(x_88, 5, x_76); -lean_ctor_set_uint8(x_88, 6, x_77); -lean_ctor_set_uint8(x_88, 7, x_78); -lean_ctor_set_uint8(x_88, 8, x_79); -lean_ctor_set_uint8(x_88, 9, x_87); -lean_ctor_set_uint8(x_88, 10, x_80); -lean_ctor_set_uint8(x_88, 11, x_81); -lean_ctor_set_uint8(x_88, 12, x_82); -lean_ctor_set_uint8(x_88, 13, x_83); -lean_ctor_set_uint8(x_88, 14, x_84); -lean_ctor_set_uint8(x_88, 15, x_85); -lean_ctor_set_uint8(x_88, 16, x_86); -x_89 = 2; -x_90 = lean_uint64_shift_right(x_37, x_89); -x_91 = lean_uint64_shift_left(x_90, x_89); -x_92 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; -x_93 = lean_uint64_lor(x_91, x_92); -x_94 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_94, 0, x_88); -lean_ctor_set(x_94, 1, x_39); -lean_ctor_set(x_94, 2, x_40); -lean_ctor_set(x_94, 3, x_41); -lean_ctor_set(x_94, 4, x_42); -lean_ctor_set(x_94, 5, x_43); -lean_ctor_set(x_94, 6, x_44); -lean_ctor_set_uint64(x_94, sizeof(void*)*7, x_93); -lean_ctor_set_uint8(x_94, sizeof(void*)*7 + 8, x_38); -lean_ctor_set_uint8(x_94, sizeof(void*)*7 + 9, x_69); -lean_ctor_set_uint8(x_94, sizeof(void*)*7 + 10, x_70); -x_95 = 0; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_58, 0); +x_69 = lean_ctor_get(x_58, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_58); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +else +{ +uint8_t x_71; uint8_t x_72; uint8_t x_73; uint8_t x_74; uint8_t x_75; uint8_t x_76; uint8_t x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; uint8_t x_82; uint8_t x_83; uint8_t x_84; uint8_t x_85; uint8_t x_86; uint8_t x_87; uint8_t x_88; uint8_t x_89; lean_object* x_90; uint64_t x_91; uint64_t x_92; uint64_t x_93; uint64_t x_94; uint64_t x_95; lean_object* x_96; uint8_t x_97; lean_object* x_98; +x_71 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 9); +x_72 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 10); +x_73 = lean_ctor_get_uint8(x_37, 0); +x_74 = lean_ctor_get_uint8(x_37, 1); +x_75 = lean_ctor_get_uint8(x_37, 2); +x_76 = lean_ctor_get_uint8(x_37, 3); +x_77 = lean_ctor_get_uint8(x_37, 4); +x_78 = lean_ctor_get_uint8(x_37, 5); +x_79 = lean_ctor_get_uint8(x_37, 6); +x_80 = lean_ctor_get_uint8(x_37, 7); +x_81 = lean_ctor_get_uint8(x_37, 8); +x_82 = lean_ctor_get_uint8(x_37, 10); +x_83 = lean_ctor_get_uint8(x_37, 11); +x_84 = lean_ctor_get_uint8(x_37, 12); +x_85 = lean_ctor_get_uint8(x_37, 13); +x_86 = lean_ctor_get_uint8(x_37, 14); +x_87 = lean_ctor_get_uint8(x_37, 15); +x_88 = lean_ctor_get_uint8(x_37, 16); +lean_dec(x_37); +x_89 = 1; +x_90 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_90, 0, x_73); +lean_ctor_set_uint8(x_90, 1, x_74); +lean_ctor_set_uint8(x_90, 2, x_75); +lean_ctor_set_uint8(x_90, 3, x_76); +lean_ctor_set_uint8(x_90, 4, x_77); +lean_ctor_set_uint8(x_90, 5, x_78); +lean_ctor_set_uint8(x_90, 6, x_79); +lean_ctor_set_uint8(x_90, 7, x_80); +lean_ctor_set_uint8(x_90, 8, x_81); +lean_ctor_set_uint8(x_90, 9, x_89); +lean_ctor_set_uint8(x_90, 10, x_82); +lean_ctor_set_uint8(x_90, 11, x_83); +lean_ctor_set_uint8(x_90, 12, x_84); +lean_ctor_set_uint8(x_90, 13, x_85); +lean_ctor_set_uint8(x_90, 14, x_86); +lean_ctor_set_uint8(x_90, 15, x_87); +lean_ctor_set_uint8(x_90, 16, x_88); +x_91 = 2; +x_92 = lean_uint64_shift_right(x_39, x_91); +x_93 = lean_uint64_shift_left(x_92, x_91); +x_94 = l_List_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonElemCore___spec__10___lambda__1___closed__9; +x_95 = lean_uint64_lor(x_93, x_94); +x_96 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_96, 0, x_90); +lean_ctor_set(x_96, 1, x_41); +lean_ctor_set(x_96, 2, x_42); +lean_ctor_set(x_96, 3, x_43); +lean_ctor_set(x_96, 4, x_44); +lean_ctor_set(x_96, 5, x_45); +lean_ctor_set(x_96, 6, x_46); +lean_ctor_set_uint64(x_96, sizeof(void*)*7, x_95); +lean_ctor_set_uint8(x_96, sizeof(void*)*7 + 8, x_40); +lean_ctor_set_uint8(x_96, sizeof(void*)*7 + 9, x_71); +lean_ctor_set_uint8(x_96, sizeof(void*)*7 + 10, x_72); +x_97 = 0; +lean_inc(x_16); +lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_31); -lean_inc(x_6); +lean_inc(x_33); +lean_inc(x_8); lean_inc(x_1); -x_96 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_6, x_31, x_95, x_10, x_94, x_12, x_13, x_14, x_36); -if (lean_obj_tag(x_96) == 0) -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = lean_ctor_get(x_97, 0); +x_98 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_8, x_33, x_97, x_12, x_96, x_14, x_15, x_16, x_38); +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_99 = lean_ctor_get(x_98, 0); lean_inc(x_99); -x_100 = lean_ctor_get(x_97, 1); +x_100 = lean_ctor_get(x_98, 1); lean_inc(x_100); -lean_dec(x_97); -x_101 = lean_unbox(x_30); -lean_dec(x_30); -x_102 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_101, x_99, x_9, x_100, x_11, x_12, x_13, x_14, x_98); -lean_dec(x_31); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_20 = x_103; -x_21 = x_104; -goto block_28; +lean_dec(x_98); +x_101 = lean_ctor_get(x_99, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_99, 1); +lean_inc(x_102); +lean_dec(x_99); +x_103 = lean_unbox(x_32); +lean_dec(x_32); +x_104 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_33, x_8, x_31, x_103, x_101, x_11, x_102, x_13, x_14, x_15, x_16, x_100); +lean_dec(x_33); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_22 = x_105; +x_23 = x_106; +goto block_30; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_33); +lean_dec(x_32); lean_dec(x_31); -lean_dec(x_30); -lean_dec(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_9); -lean_dec(x_6); +lean_dec(x_8); lean_dec(x_1); -x_105 = lean_ctor_get(x_96, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_96, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_107 = x_96; +x_107 = lean_ctor_get(x_98, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_98, 1); +lean_inc(x_108); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_109 = x_98; } else { - lean_dec_ref(x_96); - x_107 = lean_box(0); + lean_dec_ref(x_98); + x_109 = lean_box(0); } -if (lean_is_scalar(x_107)) { - x_108 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_109)) { + x_110 = lean_alloc_ctor(1, 2, 0); } else { - x_108 = x_107; + x_110 = x_109; } -lean_ctor_set(x_108, 0, x_105); -lean_ctor_set(x_108, 1, x_106); -return x_108; +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_108); +return x_110; } } } case 1: { -lean_object* x_109; lean_object* x_110; uint64_t x_111; uint8_t x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; -x_109 = lean_ctor_get(x_11, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_32, 1); -lean_inc(x_110); -lean_dec(x_32); -x_111 = lean_ctor_get_uint64(x_11, sizeof(void*)*7); -x_112 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 8); -x_113 = lean_ctor_get(x_11, 1); -lean_inc(x_113); -x_114 = lean_ctor_get(x_11, 2); -lean_inc(x_114); -x_115 = lean_ctor_get(x_11, 3); +lean_object* x_111; lean_object* x_112; uint64_t x_113; uint8_t x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_111 = lean_ctor_get(x_13, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_34, 1); +lean_inc(x_112); +lean_dec(x_34); +x_113 = lean_ctor_get_uint64(x_13, sizeof(void*)*7); +x_114 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 8); +x_115 = lean_ctor_get(x_13, 1); lean_inc(x_115); -x_116 = lean_ctor_get(x_11, 4); +x_116 = lean_ctor_get(x_13, 2); lean_inc(x_116); -x_117 = lean_ctor_get(x_11, 5); +x_117 = lean_ctor_get(x_13, 3); lean_inc(x_117); -x_118 = lean_ctor_get(x_11, 6); +x_118 = lean_ctor_get(x_13, 4); lean_inc(x_118); -x_119 = !lean_is_exclusive(x_109); -if (x_119 == 0) -{ -uint8_t x_120; uint8_t x_121; uint8_t x_122; uint64_t x_123; uint64_t x_124; uint64_t x_125; uint64_t x_126; uint64_t x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; -x_120 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 9); -x_121 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 10); -x_122 = 3; -lean_ctor_set_uint8(x_109, 9, x_122); -x_123 = 2; -x_124 = lean_uint64_shift_right(x_111, x_123); -x_125 = lean_uint64_shift_left(x_124, x_123); -x_126 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; -x_127 = lean_uint64_lor(x_125, x_126); -x_128 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_128, 0, x_109); -lean_ctor_set(x_128, 1, x_113); -lean_ctor_set(x_128, 2, x_114); -lean_ctor_set(x_128, 3, x_115); -lean_ctor_set(x_128, 4, x_116); -lean_ctor_set(x_128, 5, x_117); -lean_ctor_set(x_128, 6, x_118); -lean_ctor_set_uint64(x_128, sizeof(void*)*7, x_127); -lean_ctor_set_uint8(x_128, sizeof(void*)*7 + 8, x_112); -lean_ctor_set_uint8(x_128, sizeof(void*)*7 + 9, x_120); -lean_ctor_set_uint8(x_128, sizeof(void*)*7 + 10, x_121); -x_129 = 1; +x_119 = lean_ctor_get(x_13, 5); +lean_inc(x_119); +x_120 = lean_ctor_get(x_13, 6); +lean_inc(x_120); +x_121 = !lean_is_exclusive(x_111); +if (x_121 == 0) +{ +uint8_t x_122; uint8_t x_123; uint8_t x_124; uint64_t x_125; uint64_t x_126; uint64_t x_127; uint64_t x_128; uint64_t x_129; lean_object* x_130; uint8_t x_131; lean_object* x_132; +x_122 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 9); +x_123 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 10); +x_124 = 3; +lean_ctor_set_uint8(x_111, 9, x_124); +x_125 = 2; +x_126 = lean_uint64_shift_right(x_113, x_125); +x_127 = lean_uint64_shift_left(x_126, x_125); +x_128 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; +x_129 = lean_uint64_lor(x_127, x_128); +x_130 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_130, 0, x_111); +lean_ctor_set(x_130, 1, x_115); +lean_ctor_set(x_130, 2, x_116); +lean_ctor_set(x_130, 3, x_117); +lean_ctor_set(x_130, 4, x_118); +lean_ctor_set(x_130, 5, x_119); +lean_ctor_set(x_130, 6, x_120); +lean_ctor_set_uint64(x_130, sizeof(void*)*7, x_129); +lean_ctor_set_uint8(x_130, sizeof(void*)*7 + 8, x_114); +lean_ctor_set_uint8(x_130, sizeof(void*)*7 + 9, x_122); +lean_ctor_set_uint8(x_130, sizeof(void*)*7 + 10, x_123); +x_131 = 1; +lean_inc(x_16); +lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_31); -lean_inc(x_6); +lean_inc(x_33); +lean_inc(x_8); lean_inc(x_1); -x_130 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_6, x_31, x_129, x_10, x_128, x_12, x_13, x_14, x_110); -if (lean_obj_tag(x_130) == 0) +x_132 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_8, x_33, x_131, x_12, x_130, x_14, x_15, x_16, x_112); +if (lean_obj_tag(x_132) == 0) { -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_130, 1); -lean_inc(x_132); -lean_dec(x_130); -x_133 = lean_ctor_get(x_131, 0); +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_133 = lean_ctor_get(x_132, 0); lean_inc(x_133); -x_134 = lean_ctor_get(x_131, 1); +x_134 = lean_ctor_get(x_132, 1); lean_inc(x_134); -lean_dec(x_131); -x_135 = lean_unbox(x_30); -lean_dec(x_30); -x_136 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_135, x_133, x_9, x_134, x_11, x_12, x_13, x_14, x_132); -lean_dec(x_31); -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_136, 1); -lean_inc(x_138); -lean_dec(x_136); -x_20 = x_137; -x_21 = x_138; -goto block_28; +lean_dec(x_132); +x_135 = lean_ctor_get(x_133, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_133, 1); +lean_inc(x_136); +lean_dec(x_133); +x_137 = lean_unbox(x_32); +lean_dec(x_32); +x_138 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_33, x_8, x_31, x_137, x_135, x_11, x_136, x_13, x_14, x_15, x_16, x_134); +lean_dec(x_33); +x_139 = lean_ctor_get(x_138, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_138, 1); +lean_inc(x_140); +lean_dec(x_138); +x_22 = x_139; +x_23 = x_140; +goto block_30; } else { -uint8_t x_139; +uint8_t x_141; +lean_dec(x_33); +lean_dec(x_32); lean_dec(x_31); -lean_dec(x_30); -lean_dec(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_9); -lean_dec(x_6); +lean_dec(x_8); lean_dec(x_1); -x_139 = !lean_is_exclusive(x_130); -if (x_139 == 0) +x_141 = !lean_is_exclusive(x_132); +if (x_141 == 0) { -return x_130; +return x_132; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_130, 0); -x_141 = lean_ctor_get(x_130, 1); -lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_130); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -return x_142; -} -} -} -else -{ -uint8_t x_143; uint8_t x_144; uint8_t x_145; uint8_t x_146; uint8_t x_147; uint8_t x_148; uint8_t x_149; uint8_t x_150; uint8_t x_151; uint8_t x_152; uint8_t x_153; uint8_t x_154; uint8_t x_155; uint8_t x_156; uint8_t x_157; uint8_t x_158; uint8_t x_159; uint8_t x_160; uint8_t x_161; lean_object* x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; uint64_t x_166; uint64_t x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; -x_143 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 9); -x_144 = lean_ctor_get_uint8(x_11, sizeof(void*)*7 + 10); -x_145 = lean_ctor_get_uint8(x_109, 0); -x_146 = lean_ctor_get_uint8(x_109, 1); -x_147 = lean_ctor_get_uint8(x_109, 2); -x_148 = lean_ctor_get_uint8(x_109, 3); -x_149 = lean_ctor_get_uint8(x_109, 4); -x_150 = lean_ctor_get_uint8(x_109, 5); -x_151 = lean_ctor_get_uint8(x_109, 6); -x_152 = lean_ctor_get_uint8(x_109, 7); -x_153 = lean_ctor_get_uint8(x_109, 8); -x_154 = lean_ctor_get_uint8(x_109, 10); -x_155 = lean_ctor_get_uint8(x_109, 11); -x_156 = lean_ctor_get_uint8(x_109, 12); -x_157 = lean_ctor_get_uint8(x_109, 13); -x_158 = lean_ctor_get_uint8(x_109, 14); -x_159 = lean_ctor_get_uint8(x_109, 15); -x_160 = lean_ctor_get_uint8(x_109, 16); -lean_dec(x_109); -x_161 = 3; -x_162 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_162, 0, x_145); -lean_ctor_set_uint8(x_162, 1, x_146); -lean_ctor_set_uint8(x_162, 2, x_147); -lean_ctor_set_uint8(x_162, 3, x_148); -lean_ctor_set_uint8(x_162, 4, x_149); -lean_ctor_set_uint8(x_162, 5, x_150); -lean_ctor_set_uint8(x_162, 6, x_151); -lean_ctor_set_uint8(x_162, 7, x_152); -lean_ctor_set_uint8(x_162, 8, x_153); -lean_ctor_set_uint8(x_162, 9, x_161); -lean_ctor_set_uint8(x_162, 10, x_154); -lean_ctor_set_uint8(x_162, 11, x_155); -lean_ctor_set_uint8(x_162, 12, x_156); -lean_ctor_set_uint8(x_162, 13, x_157); -lean_ctor_set_uint8(x_162, 14, x_158); -lean_ctor_set_uint8(x_162, 15, x_159); -lean_ctor_set_uint8(x_162, 16, x_160); -x_163 = 2; -x_164 = lean_uint64_shift_right(x_111, x_163); -x_165 = lean_uint64_shift_left(x_164, x_163); -x_166 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; -x_167 = lean_uint64_lor(x_165, x_166); -x_168 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_168, 0, x_162); -lean_ctor_set(x_168, 1, x_113); -lean_ctor_set(x_168, 2, x_114); -lean_ctor_set(x_168, 3, x_115); -lean_ctor_set(x_168, 4, x_116); -lean_ctor_set(x_168, 5, x_117); -lean_ctor_set(x_168, 6, x_118); -lean_ctor_set_uint64(x_168, sizeof(void*)*7, x_167); -lean_ctor_set_uint8(x_168, sizeof(void*)*7 + 8, x_112); -lean_ctor_set_uint8(x_168, sizeof(void*)*7 + 9, x_143); -lean_ctor_set_uint8(x_168, sizeof(void*)*7 + 10, x_144); -x_169 = 1; +lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_142 = lean_ctor_get(x_132, 0); +x_143 = lean_ctor_get(x_132, 1); +lean_inc(x_143); +lean_inc(x_142); +lean_dec(x_132); +x_144 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +return x_144; +} +} +} +else +{ +uint8_t x_145; uint8_t x_146; uint8_t x_147; uint8_t x_148; uint8_t x_149; uint8_t x_150; uint8_t x_151; uint8_t x_152; uint8_t x_153; uint8_t x_154; uint8_t x_155; uint8_t x_156; uint8_t x_157; uint8_t x_158; uint8_t x_159; uint8_t x_160; uint8_t x_161; uint8_t x_162; uint8_t x_163; lean_object* x_164; uint64_t x_165; uint64_t x_166; uint64_t x_167; uint64_t x_168; uint64_t x_169; lean_object* x_170; uint8_t x_171; lean_object* x_172; +x_145 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 9); +x_146 = lean_ctor_get_uint8(x_13, sizeof(void*)*7 + 10); +x_147 = lean_ctor_get_uint8(x_111, 0); +x_148 = lean_ctor_get_uint8(x_111, 1); +x_149 = lean_ctor_get_uint8(x_111, 2); +x_150 = lean_ctor_get_uint8(x_111, 3); +x_151 = lean_ctor_get_uint8(x_111, 4); +x_152 = lean_ctor_get_uint8(x_111, 5); +x_153 = lean_ctor_get_uint8(x_111, 6); +x_154 = lean_ctor_get_uint8(x_111, 7); +x_155 = lean_ctor_get_uint8(x_111, 8); +x_156 = lean_ctor_get_uint8(x_111, 10); +x_157 = lean_ctor_get_uint8(x_111, 11); +x_158 = lean_ctor_get_uint8(x_111, 12); +x_159 = lean_ctor_get_uint8(x_111, 13); +x_160 = lean_ctor_get_uint8(x_111, 14); +x_161 = lean_ctor_get_uint8(x_111, 15); +x_162 = lean_ctor_get_uint8(x_111, 16); +lean_dec(x_111); +x_163 = 3; +x_164 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_164, 0, x_147); +lean_ctor_set_uint8(x_164, 1, x_148); +lean_ctor_set_uint8(x_164, 2, x_149); +lean_ctor_set_uint8(x_164, 3, x_150); +lean_ctor_set_uint8(x_164, 4, x_151); +lean_ctor_set_uint8(x_164, 5, x_152); +lean_ctor_set_uint8(x_164, 6, x_153); +lean_ctor_set_uint8(x_164, 7, x_154); +lean_ctor_set_uint8(x_164, 8, x_155); +lean_ctor_set_uint8(x_164, 9, x_163); +lean_ctor_set_uint8(x_164, 10, x_156); +lean_ctor_set_uint8(x_164, 11, x_157); +lean_ctor_set_uint8(x_164, 12, x_158); +lean_ctor_set_uint8(x_164, 13, x_159); +lean_ctor_set_uint8(x_164, 14, x_160); +lean_ctor_set_uint8(x_164, 15, x_161); +lean_ctor_set_uint8(x_164, 16, x_162); +x_165 = 2; +x_166 = lean_uint64_shift_right(x_113, x_165); +x_167 = lean_uint64_shift_left(x_166, x_165); +x_168 = l_Lean_Meta_Grind_Canon_canonInst___closed__1; +x_169 = lean_uint64_lor(x_167, x_168); +x_170 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_170, 0, x_164); +lean_ctor_set(x_170, 1, x_115); +lean_ctor_set(x_170, 2, x_116); +lean_ctor_set(x_170, 3, x_117); +lean_ctor_set(x_170, 4, x_118); +lean_ctor_set(x_170, 5, x_119); +lean_ctor_set(x_170, 6, x_120); +lean_ctor_set_uint64(x_170, sizeof(void*)*7, x_169); +lean_ctor_set_uint8(x_170, sizeof(void*)*7 + 8, x_114); +lean_ctor_set_uint8(x_170, sizeof(void*)*7 + 9, x_145); +lean_ctor_set_uint8(x_170, sizeof(void*)*7 + 10, x_146); +x_171 = 1; +lean_inc(x_16); +lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_31); -lean_inc(x_6); +lean_inc(x_33); +lean_inc(x_8); lean_inc(x_1); -x_170 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_6, x_31, x_169, x_10, x_168, x_12, x_13, x_14, x_110); -if (lean_obj_tag(x_170) == 0) +x_172 = l_Lean_Meta_Grind_Canon_canonElemCore(x_1, x_8, x_33, x_171, x_12, x_170, x_14, x_15, x_16, x_112); +if (lean_obj_tag(x_172) == 0) { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_171 = lean_ctor_get(x_170, 0); -lean_inc(x_171); -x_172 = lean_ctor_get(x_170, 1); -lean_inc(x_172); -lean_dec(x_170); -x_173 = lean_ctor_get(x_171, 0); +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; uint8_t x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_173 = lean_ctor_get(x_172, 0); lean_inc(x_173); -x_174 = lean_ctor_get(x_171, 1); +x_174 = lean_ctor_get(x_172, 1); lean_inc(x_174); -lean_dec(x_171); -x_175 = lean_unbox(x_30); -lean_dec(x_30); -x_176 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_175, x_173, x_9, x_174, x_11, x_12, x_13, x_14, x_172); -lean_dec(x_31); -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); -lean_inc(x_178); -lean_dec(x_176); -x_20 = x_177; -x_21 = x_178; -goto block_28; +lean_dec(x_172); +x_175 = lean_ctor_get(x_173, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_173, 1); +lean_inc(x_176); +lean_dec(x_173); +x_177 = lean_unbox(x_32); +lean_dec(x_32); +x_178 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_33, x_8, x_31, x_177, x_175, x_11, x_176, x_13, x_14, x_15, x_16, x_174); +lean_dec(x_33); +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_178, 1); +lean_inc(x_180); +lean_dec(x_178); +x_22 = x_179; +x_23 = x_180; +goto block_30; } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_33); +lean_dec(x_32); lean_dec(x_31); -lean_dec(x_30); -lean_dec(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_9); -lean_dec(x_6); +lean_dec(x_8); lean_dec(x_1); -x_179 = lean_ctor_get(x_170, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_170, 1); -lean_inc(x_180); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - lean_ctor_release(x_170, 1); - x_181 = x_170; +x_181 = lean_ctor_get(x_172, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_172, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + x_183 = x_172; } else { - lean_dec_ref(x_170); - x_181 = lean_box(0); + lean_dec_ref(x_172); + x_183 = lean_box(0); } -if (lean_is_scalar(x_181)) { - x_182 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_183)) { + x_184 = lean_alloc_ctor(1, 2, 0); } else { - x_182 = x_181; + x_184 = x_183; } -lean_ctor_set(x_182, 0, x_179); -lean_ctor_set(x_182, 1, x_180); -return x_182; +lean_ctor_set(x_184, 0, x_181); +lean_ctor_set(x_184, 1, x_182); +return x_184; } } } default: { -lean_object* x_183; lean_object* x_184; -x_183 = lean_ctor_get(x_32, 1); -lean_inc(x_183); -lean_dec(x_32); +lean_object* x_185; lean_object* x_186; +x_185 = lean_ctor_get(x_34, 1); +lean_inc(x_185); +lean_dec(x_34); +lean_inc(x_16); +lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); -lean_inc(x_12); lean_inc(x_11); -lean_inc(x_9); -lean_inc(x_31); -x_184 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_31, x_9, x_10, x_11, x_12, x_13, x_14, x_183); -if (lean_obj_tag(x_184) == 0) +lean_inc(x_33); +x_186 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_33, x_11, x_12, x_13, x_14, x_15, x_16, x_185); +if (lean_obj_tag(x_186) == 0) { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_185 = lean_ctor_get(x_184, 0); -lean_inc(x_185); -x_186 = lean_ctor_get(x_184, 1); -lean_inc(x_186); -lean_dec(x_184); -x_187 = lean_ctor_get(x_185, 0); +lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_187 = lean_ctor_get(x_186, 0); lean_inc(x_187); -x_188 = lean_ctor_get(x_185, 1); +x_188 = lean_ctor_get(x_186, 1); lean_inc(x_188); -lean_dec(x_185); -x_189 = lean_unbox(x_30); -lean_dec(x_30); -x_190 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_31, x_6, x_29, x_189, x_187, x_9, x_188, x_11, x_12, x_13, x_14, x_186); -lean_dec(x_31); -x_191 = lean_ctor_get(x_190, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_190, 1); -lean_inc(x_192); -lean_dec(x_190); -x_20 = x_191; -x_21 = x_192; -goto block_28; +lean_dec(x_186); +x_189 = lean_ctor_get(x_187, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_187, 1); +lean_inc(x_190); +lean_dec(x_187); +x_191 = lean_unbox(x_32); +lean_dec(x_32); +x_192 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___lambda__1(x_33, x_8, x_31, x_191, x_189, x_11, x_190, x_13, x_14, x_15, x_16, x_188); +lean_dec(x_33); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_22 = x_193; +x_23 = x_194; +goto block_30; } else { -uint8_t x_193; +uint8_t x_195; +lean_dec(x_33); +lean_dec(x_32); lean_dec(x_31); -lean_dec(x_30); -lean_dec(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_9); -lean_dec(x_6); +lean_dec(x_8); lean_dec(x_1); -x_193 = !lean_is_exclusive(x_184); -if (x_193 == 0) +x_195 = !lean_is_exclusive(x_186); +if (x_195 == 0) { -return x_184; +return x_186; } else { -lean_object* x_194; lean_object* x_195; lean_object* x_196; -x_194 = lean_ctor_get(x_184, 0); -x_195 = lean_ctor_get(x_184, 1); -lean_inc(x_195); -lean_inc(x_194); -lean_dec(x_184); -x_196 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -return x_196; +lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_196 = lean_ctor_get(x_186, 0); +x_197 = lean_ctor_get(x_186, 1); +lean_inc(x_197); +lean_inc(x_196); +lean_dec(x_186); +x_198 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_198, 0, x_196); +lean_ctor_set(x_198, 1, x_197); +return x_198; } } } @@ -6297,866 +6264,694 @@ return x_196; } else { -uint8_t x_197; +uint8_t x_199; +lean_dec(x_33); +lean_dec(x_32); lean_dec(x_31); -lean_dec(x_30); -lean_dec(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_6); +lean_dec(x_8); lean_dec(x_1); -x_197 = !lean_is_exclusive(x_32); -if (x_197 == 0) +x_199 = !lean_is_exclusive(x_34); +if (x_199 == 0) { -return x_32; +return x_34; } else { -lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_198 = lean_ctor_get(x_32, 0); -x_199 = lean_ctor_get(x_32, 1); -lean_inc(x_199); -lean_inc(x_198); -lean_dec(x_32); -x_200 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_200, 0, x_198); -lean_ctor_set(x_200, 1, x_199); -return x_200; +lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_200 = lean_ctor_get(x_34, 0); +x_201 = lean_ctor_get(x_34, 1); +lean_inc(x_201); +lean_inc(x_200); +lean_dec(x_34); +x_202 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_202, 0, x_200); +lean_ctor_set(x_202, 1, x_201); +return x_202; } } +block_30: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_ctor_get(x_6, 2); +x_28 = lean_nat_add(x_8, x_27); +lean_dec(x_8); +x_7 = x_26; +x_8 = x_28; +x_9 = lean_box(0); +x_10 = lean_box(0); +x_12 = x_25; +x_17 = x_23; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1() { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_ref_take(x_3, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2() { +_start: { -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_10, 0); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3() { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; uint8_t x_34; -x_14 = lean_ctor_get(x_10, 1); -x_15 = lean_ctor_get(x_12, 0); -x_16 = lean_ctor_get(x_12, 1); -x_17 = lean_array_get_size(x_16); -x_18 = lean_ptr_addr(x_1); -x_19 = lean_usize_to_uint64(x_18); -x_20 = 11; -x_21 = lean_uint64_mix_hash(x_19, x_20); -x_22 = 32; -x_23 = lean_uint64_shift_right(x_21, x_22); -x_24 = lean_uint64_xor(x_21, x_23); -x_25 = 16; -x_26 = lean_uint64_shift_right(x_24, x_25); -x_27 = lean_uint64_xor(x_24, x_26); -x_28 = lean_uint64_to_usize(x_27); -x_29 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_30 = 1; -x_31 = lean_usize_sub(x_29, x_30); -x_32 = lean_usize_land(x_28, x_31); -x_33 = lean_array_uget(x_16, x_32); -x_34 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_1, x_33); -if (x_34 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4() { +_start: { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_35 = lean_unsigned_to_nat(1u); -x_36 = lean_nat_add(x_15, x_35); -lean_dec(x_15); -lean_inc(x_2); -x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_1); -lean_ctor_set(x_37, 1, x_2); -lean_ctor_set(x_37, 2, x_33); -x_38 = lean_array_uset(x_16, x_32, x_37); -x_39 = lean_unsigned_to_nat(4u); -x_40 = lean_nat_mul(x_36, x_39); -x_41 = lean_unsigned_to_nat(3u); -x_42 = lean_nat_div(x_40, x_41); -lean_dec(x_40); -x_43 = lean_array_get_size(x_38); -x_44 = lean_nat_dec_le(x_42, x_43); -lean_dec(x_43); -lean_dec(x_42); -if (x_44 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1; +x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2; +x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(x_38); -lean_ctor_set(x_12, 1, x_45); -lean_ctor_set(x_12, 0, x_36); -x_46 = lean_st_ref_set(x_3, x_12, x_14); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +switch (lean_obj_tag(x_2)) { +case 0: { -lean_object* x_48; -x_48 = lean_ctor_get(x_46, 0); -lean_dec(x_48); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_46, 0, x_10); -return x_46; +lean_object* x_12; lean_object* x_91; lean_object* x_145; uint8_t x_146; +lean_dec(x_4); +x_145 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_146 = l_Lean_Expr_isConstOf(x_2, x_145); +if (x_146 == 0) +{ +lean_object* x_147; +x_147 = lean_box(0); +x_91 = x_147; +goto block_144; } else { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_10); -lean_ctor_set(x_50, 1, x_49); -return x_50; -} +lean_object* x_148; lean_object* x_149; uint8_t x_150; +x_148 = lean_array_get_size(x_3); +x_149 = lean_unsigned_to_nat(2u); +x_150 = lean_nat_dec_eq(x_148, x_149); +if (x_150 == 0) +{ +lean_object* x_151; +lean_dec(x_148); +x_151 = lean_box(0); +x_91 = x_151; +goto block_144; } else { -lean_object* x_51; uint8_t x_52; -lean_ctor_set(x_12, 1, x_38); -lean_ctor_set(x_12, 0, x_36); -x_51 = lean_st_ref_set(x_3, x_12, x_14); -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) +lean_object* x_152; uint8_t x_153; +x_152 = lean_unsigned_to_nat(0u); +x_153 = lean_nat_dec_lt(x_152, x_148); +lean_dec(x_148); +if (x_153 == 0) { -lean_object* x_53; -x_53 = lean_ctor_get(x_51, 0); -lean_dec(x_53); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_51, 0, x_10); -return x_51; +lean_object* x_154; lean_object* x_155; +x_154 = l_Lean_instInhabitedExpr; +x_155 = l_outOfBounds___rarg(x_154); +x_12 = x_155; +goto block_90; } else { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); -lean_dec(x_51); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_10); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_object* x_156; +x_156 = lean_array_fget(x_3, x_152); +x_12 = x_156; +goto block_90; } } } -else +block_90: { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_56 = lean_box(0); -x_57 = lean_array_uset(x_16, x_32, x_56); -lean_inc(x_2); -x_58 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_33); -x_59 = lean_array_uset(x_57, x_32, x_58); -lean_ctor_set(x_12, 1, x_59); -x_60 = lean_st_ref_set(x_3, x_12, x_14); -x_61 = !lean_is_exclusive(x_60); -if (x_61 == 0) +lean_object* x_13; +lean_inc(x_12); +x_13 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_62; -x_62 = lean_ctor_get(x_60, 0); -lean_dec(x_62); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_60, 0, x_10); -return x_60; +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = lean_ctor_get(x_18, 2); +lean_inc(x_19); +lean_inc(x_19); +x_20 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_19, x_17); +if (lean_obj_tag(x_20) == 0) +{ +size_t x_21; size_t x_22; uint8_t x_23; +x_21 = lean_ptr_addr(x_12); +lean_dec(x_12); +x_22 = lean_ptr_addr(x_17); +x_23 = lean_usize_dec_eq(x_21, x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_1); +x_24 = lean_unsigned_to_nat(0u); +lean_inc(x_17); +x_25 = lean_array_set(x_3, x_24, x_17); +x_26 = l_Lean_mkAppN(x_2, x_25); +lean_dec(x_25); +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); +lean_inc(x_26); +x_29 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_19, x_17, x_26); +x_30 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_28); +lean_ctor_set(x_30, 2, x_29); +lean_ctor_set(x_15, 1, x_30); +lean_ctor_set(x_15, 0, x_26); +return x_13; } else { -lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_60, 1); -lean_inc(x_63); -lean_dec(x_60); -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_10); -lean_ctor_set(x_64, 1, x_63); -return x_64; -} +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_3); +lean_dec(x_2); +x_31 = lean_ctor_get(x_18, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_18, 1); +lean_inc(x_32); +lean_dec(x_18); +lean_inc(x_1); +x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_19, x_17, x_1); +x_34 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_33); +lean_ctor_set(x_15, 1, x_34); +lean_ctor_set(x_15, 0, x_1); +return x_13; } } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; size_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; uint64_t x_76; uint64_t x_77; uint64_t x_78; size_t x_79; size_t x_80; size_t x_81; size_t x_82; size_t x_83; lean_object* x_84; uint8_t x_85; -x_65 = lean_ctor_get(x_10, 1); -x_66 = lean_ctor_get(x_12, 0); -x_67 = lean_ctor_get(x_12, 1); -lean_inc(x_67); -lean_inc(x_66); +lean_object* x_35; +lean_dec(x_19); +lean_dec(x_17); lean_dec(x_12); -x_68 = lean_array_get_size(x_67); -x_69 = lean_ptr_addr(x_1); -x_70 = lean_usize_to_uint64(x_69); -x_71 = 11; -x_72 = lean_uint64_mix_hash(x_70, x_71); -x_73 = 32; -x_74 = lean_uint64_shift_right(x_72, x_73); -x_75 = lean_uint64_xor(x_72, x_74); -x_76 = 16; -x_77 = lean_uint64_shift_right(x_75, x_76); -x_78 = lean_uint64_xor(x_75, x_77); -x_79 = lean_uint64_to_usize(x_78); -x_80 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_81 = 1; -x_82 = lean_usize_sub(x_80, x_81); -x_83 = lean_usize_land(x_79, x_82); -x_84 = lean_array_uget(x_67, x_83); -x_85 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_1, x_84); -if (x_85 == 0) +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_35 = lean_ctor_get(x_20, 0); +lean_inc(x_35); +lean_dec(x_20); +lean_ctor_set(x_15, 0, x_35); +return x_13; +} +} +else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_86 = lean_unsigned_to_nat(1u); -x_87 = lean_nat_add(x_66, x_86); -lean_dec(x_66); -lean_inc(x_2); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_1); -lean_ctor_set(x_88, 1, x_2); -lean_ctor_set(x_88, 2, x_84); -x_89 = lean_array_uset(x_67, x_83, x_88); -x_90 = lean_unsigned_to_nat(4u); -x_91 = lean_nat_mul(x_87, x_90); -x_92 = lean_unsigned_to_nat(3u); -x_93 = lean_nat_div(x_91, x_92); -lean_dec(x_91); -x_94 = lean_array_get_size(x_89); -x_95 = lean_nat_dec_le(x_93, x_94); -lean_dec(x_94); -lean_dec(x_93); -if (x_95 == 0) +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_15, 0); +x_37 = lean_ctor_get(x_15, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_15); +x_38 = lean_ctor_get(x_37, 2); +lean_inc(x_38); +lean_inc(x_38); +x_39 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_38, x_36); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_96 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(x_89); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_87); -lean_ctor_set(x_97, 1, x_96); -x_98 = lean_st_ref_set(x_3, x_97, x_65); -x_99 = lean_ctor_get(x_98, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_100 = x_98; -} else { - lean_dec_ref(x_98); - x_100 = lean_box(0); +size_t x_40; size_t x_41; uint8_t x_42; +x_40 = lean_ptr_addr(x_12); +lean_dec(x_12); +x_41 = lean_ptr_addr(x_36); +x_42 = lean_usize_dec_eq(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_1); +x_43 = lean_unsigned_to_nat(0u); +lean_inc(x_36); +x_44 = lean_array_set(x_3, x_43, x_36); +x_45 = l_Lean_mkAppN(x_2, x_44); +lean_dec(x_44); +x_46 = lean_ctor_get(x_37, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_37, 1); +lean_inc(x_47); +lean_dec(x_37); +lean_inc(x_45); +x_48 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_38, x_36, x_45); +x_49 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_49, 0, x_46); +lean_ctor_set(x_49, 1, x_47); +lean_ctor_set(x_49, 2, x_48); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_45); +lean_ctor_set(x_50, 1, x_49); +lean_ctor_set(x_13, 0, x_50); +return x_13; } -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -if (lean_is_scalar(x_100)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_100; +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_3); +lean_dec(x_2); +x_51 = lean_ctor_get(x_37, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_37, 1); +lean_inc(x_52); +lean_dec(x_37); +lean_inc(x_1); +x_53 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_38, x_36, x_1); +x_54 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_54, 0, x_51); +lean_ctor_set(x_54, 1, x_52); +lean_ctor_set(x_54, 2, x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_1); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_13, 0, x_55); +return x_13; } -lean_ctor_set(x_101, 0, x_10); -lean_ctor_set(x_101, 1, x_99); -return x_101; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_87); -lean_ctor_set(x_102, 1, x_89); -x_103 = lean_st_ref_set(x_3, x_102, x_65); -x_104 = lean_ctor_get(x_103, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_105 = x_103; -} else { - lean_dec_ref(x_103); - x_105 = lean_box(0); -} -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(0, 2, 0); -} else { - x_106 = x_105; +lean_object* x_56; lean_object* x_57; +lean_dec(x_38); +lean_dec(x_36); +lean_dec(x_12); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_56 = lean_ctor_get(x_39, 0); +lean_inc(x_56); +lean_dec(x_39); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_37); +lean_ctor_set(x_13, 0, x_57); +return x_13; } -lean_ctor_set(x_106, 0, x_10); -lean_ctor_set(x_106, 1, x_104); -return x_106; } } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_107 = lean_box(0); -x_108 = lean_array_uset(x_67, x_83, x_107); -lean_inc(x_2); -x_109 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_84); -x_110 = lean_array_uset(x_108, x_83, x_109); -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_66); -lean_ctor_set(x_111, 1, x_110); -x_112 = lean_st_ref_set(x_3, x_111, x_65); -x_113 = lean_ctor_get(x_112, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_114 = x_112; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_58 = lean_ctor_get(x_13, 0); +x_59 = lean_ctor_get(x_13, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_13); +x_60 = lean_ctor_get(x_58, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_62 = x_58; } else { - lean_dec_ref(x_112); - x_114 = lean_box(0); + lean_dec_ref(x_58); + x_62 = lean_box(0); } -lean_ctor_set(x_10, 1, x_4); -lean_ctor_set(x_10, 0, x_2); -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(0, 2, 0); +x_63 = lean_ctor_get(x_61, 2); +lean_inc(x_63); +lean_inc(x_63); +x_64 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_63, x_60); +if (lean_obj_tag(x_64) == 0) +{ +size_t x_65; size_t x_66; uint8_t x_67; +x_65 = lean_ptr_addr(x_12); +lean_dec(x_12); +x_66 = lean_ptr_addr(x_60); +x_67 = lean_usize_dec_eq(x_65, x_66); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_1); +x_68 = lean_unsigned_to_nat(0u); +lean_inc(x_60); +x_69 = lean_array_set(x_3, x_68, x_60); +x_70 = l_Lean_mkAppN(x_2, x_69); +lean_dec(x_69); +x_71 = lean_ctor_get(x_61, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_61, 1); +lean_inc(x_72); +lean_dec(x_61); +lean_inc(x_70); +x_73 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_63, x_60, x_70); +x_74 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +lean_ctor_set(x_74, 2, x_73); +if (lean_is_scalar(x_62)) { + x_75 = lean_alloc_ctor(0, 2, 0); } else { - x_115 = x_114; -} -lean_ctor_set(x_115, 0, x_10); -lean_ctor_set(x_115, 1, x_113); -return x_115; -} + x_75 = x_62; } +lean_ctor_set(x_75, 0, x_70); +lean_ctor_set(x_75, 1, x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_59); +return x_76; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; size_t x_122; uint64_t x_123; uint64_t x_124; uint64_t x_125; uint64_t x_126; uint64_t x_127; uint64_t x_128; uint64_t x_129; uint64_t x_130; uint64_t x_131; size_t x_132; size_t x_133; size_t x_134; size_t x_135; size_t x_136; lean_object* x_137; uint8_t x_138; -x_116 = lean_ctor_get(x_10, 0); -x_117 = lean_ctor_get(x_10, 1); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_10); -x_118 = lean_ctor_get(x_116, 0); -lean_inc(x_118); -x_119 = lean_ctor_get(x_116, 1); -lean_inc(x_119); -if (lean_is_exclusive(x_116)) { - lean_ctor_release(x_116, 0); - lean_ctor_release(x_116, 1); - x_120 = x_116; -} else { - lean_dec_ref(x_116); - x_120 = lean_box(0); -} -x_121 = lean_array_get_size(x_119); -x_122 = lean_ptr_addr(x_1); -x_123 = lean_usize_to_uint64(x_122); -x_124 = 11; -x_125 = lean_uint64_mix_hash(x_123, x_124); -x_126 = 32; -x_127 = lean_uint64_shift_right(x_125, x_126); -x_128 = lean_uint64_xor(x_125, x_127); -x_129 = 16; -x_130 = lean_uint64_shift_right(x_128, x_129); -x_131 = lean_uint64_xor(x_128, x_130); -x_132 = lean_uint64_to_usize(x_131); -x_133 = lean_usize_of_nat(x_121); -lean_dec(x_121); -x_134 = 1; -x_135 = lean_usize_sub(x_133, x_134); -x_136 = lean_usize_land(x_132, x_135); -x_137 = lean_array_uget(x_119, x_136); -x_138 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_1, x_137); -if (x_138 == 0) -{ -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; -x_139 = lean_unsigned_to_nat(1u); -x_140 = lean_nat_add(x_118, x_139); -lean_dec(x_118); -lean_inc(x_2); -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_1); -lean_ctor_set(x_141, 1, x_2); -lean_ctor_set(x_141, 2, x_137); -x_142 = lean_array_uset(x_119, x_136, x_141); -x_143 = lean_unsigned_to_nat(4u); -x_144 = lean_nat_mul(x_140, x_143); -x_145 = lean_unsigned_to_nat(3u); -x_146 = lean_nat_div(x_144, x_145); -lean_dec(x_144); -x_147 = lean_array_get_size(x_142); -x_148 = lean_nat_dec_le(x_146, x_147); -lean_dec(x_147); -lean_dec(x_146); -if (x_148 == 0) -{ -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_149 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__3(x_142); -if (lean_is_scalar(x_120)) { - x_150 = lean_alloc_ctor(0, 2, 0); -} else { - x_150 = x_120; -} -lean_ctor_set(x_150, 0, x_140); -lean_ctor_set(x_150, 1, x_149); -x_151 = lean_st_ref_set(x_3, x_150, x_117); -x_152 = lean_ctor_get(x_151, 1); -lean_inc(x_152); -if (lean_is_exclusive(x_151)) { - lean_ctor_release(x_151, 0); - lean_ctor_release(x_151, 1); - x_153 = x_151; +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_dec(x_3); +lean_dec(x_2); +x_77 = lean_ctor_get(x_61, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_61, 1); +lean_inc(x_78); +lean_dec(x_61); +lean_inc(x_1); +x_79 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_63, x_60, x_1); +x_80 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_78); +lean_ctor_set(x_80, 2, x_79); +if (lean_is_scalar(x_62)) { + x_81 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_151); - x_153 = lean_box(0); + x_81 = x_62; } -x_154 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_154, 0, x_2); -lean_ctor_set(x_154, 1, x_4); -if (lean_is_scalar(x_153)) { - x_155 = lean_alloc_ctor(0, 2, 0); -} else { - x_155 = x_153; +lean_ctor_set(x_81, 0, x_1); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_59); +return x_82; } -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_152); -return x_155; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -if (lean_is_scalar(x_120)) { - x_156 = lean_alloc_ctor(0, 2, 0); -} else { - x_156 = x_120; -} -lean_ctor_set(x_156, 0, x_140); -lean_ctor_set(x_156, 1, x_142); -x_157 = lean_st_ref_set(x_3, x_156, x_117); -x_158 = lean_ctor_get(x_157, 1); -lean_inc(x_158); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_159 = x_157; +lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_63); +lean_dec(x_60); +lean_dec(x_12); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_83 = lean_ctor_get(x_64, 0); +lean_inc(x_83); +lean_dec(x_64); +if (lean_is_scalar(x_62)) { + x_84 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_157); - x_159 = lean_box(0); + x_84 = x_62; } -x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_2); -lean_ctor_set(x_160, 1, x_4); -if (lean_is_scalar(x_159)) { - x_161 = lean_alloc_ctor(0, 2, 0); -} else { - x_161 = x_159; +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_61); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_59); +return x_85; } -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_158); -return x_161; } } else { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_162 = lean_box(0); -x_163 = lean_array_uset(x_119, x_136, x_162); -lean_inc(x_2); -x_164 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_1, x_2, x_137); -x_165 = lean_array_uset(x_163, x_136, x_164); -if (lean_is_scalar(x_120)) { - x_166 = lean_alloc_ctor(0, 2, 0); -} else { - x_166 = x_120; -} -lean_ctor_set(x_166, 0, x_118); -lean_ctor_set(x_166, 1, x_165); -x_167 = lean_st_ref_set(x_3, x_166, x_117); -x_168 = lean_ctor_get(x_167, 1); -lean_inc(x_168); -if (lean_is_exclusive(x_167)) { - lean_ctor_release(x_167, 0); - lean_ctor_release(x_167, 1); - x_169 = x_167; -} else { - lean_dec_ref(x_167); - x_169 = lean_box(0); -} -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_2); -lean_ctor_set(x_170, 1, x_4); -if (lean_is_scalar(x_169)) { - x_171 = lean_alloc_ctor(0, 2, 0); -} else { - x_171 = x_169; -} -lean_ctor_set(x_171, 0, x_170); -lean_ctor_set(x_171, 1, x_168); -return x_171; -} -} -} -} -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1() { -_start: +uint8_t x_86; +lean_dec(x_12); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_86 = !lean_is_exclusive(x_13); +if (x_86 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} +return x_13; } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; -} +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_13, 0); +x_88 = lean_ctor_get(x_13, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_13); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); -return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4() { -_start: +block_144: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1; -x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2; -x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +lean_object* x_92; +lean_dec(x_91); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_92 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_92) == 0) { -switch (lean_obj_tag(x_2)) { -case 0: +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_93, 0); +lean_inc(x_95); +lean_dec(x_93); +x_96 = lean_array_get_size(x_3); +x_97 = lean_unsigned_to_nat(0u); +x_98 = lean_unsigned_to_nat(1u); +lean_inc(x_96); +x_99 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_96); +lean_ctor_set(x_99, 2, x_98); +x_100 = 0; +x_101 = lean_box(x_100); +lean_inc(x_3); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_3); +lean_ctor_set(x_102, 1, x_101); +lean_inc(x_2); +x_103 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_95, x_96, x_99, x_99, x_102, x_97, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_94); +lean_dec(x_99); +lean_dec(x_96); +lean_dec(x_95); +lean_dec(x_3); +if (lean_obj_tag(x_103) == 0) { -lean_object* x_12; lean_object* x_43; lean_object* x_77; uint8_t x_78; -lean_dec(x_4); -x_77 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_78 = l_Lean_Expr_isConstOf(x_2, x_77); -if (x_78 == 0) +lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +x_107 = lean_unbox(x_106); +lean_dec(x_106); +if (x_107 == 0) { -lean_object* x_79; -x_79 = lean_box(0); -x_43 = x_79; -goto block_76; -} -else +uint8_t x_108; +lean_dec(x_105); +lean_dec(x_2); +x_108 = !lean_is_exclusive(x_103); +if (x_108 == 0) { -lean_object* x_80; lean_object* x_81; uint8_t x_82; -x_80 = lean_array_get_size(x_3); -x_81 = lean_unsigned_to_nat(2u); -x_82 = lean_nat_dec_eq(x_80, x_81); -if (x_82 == 0) +lean_object* x_109; uint8_t x_110; +x_109 = lean_ctor_get(x_103, 0); +lean_dec(x_109); +x_110 = !lean_is_exclusive(x_104); +if (x_110 == 0) { -lean_object* x_83; -lean_dec(x_80); -x_83 = lean_box(0); -x_43 = x_83; -goto block_76; +lean_object* x_111; +x_111 = lean_ctor_get(x_104, 0); +lean_dec(x_111); +lean_ctor_set(x_104, 0, x_1); +return x_103; } else { -lean_object* x_84; uint8_t x_85; -x_84 = lean_unsigned_to_nat(0u); -x_85 = lean_nat_dec_lt(x_84, x_80); -lean_dec(x_80); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; -x_86 = l_Lean_instInhabitedExpr; -x_87 = l_outOfBounds___rarg(x_86); -x_12 = x_87; -goto block_42; +lean_object* x_112; lean_object* x_113; +x_112 = lean_ctor_get(x_104, 1); +lean_inc(x_112); +lean_dec(x_104); +x_113 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_113, 0, x_1); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_103, 0, x_113); +return x_103; +} } else { -lean_object* x_88; -x_88 = lean_array_fget(x_3, x_84); -x_12 = x_88; -goto block_42; +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_114 = lean_ctor_get(x_103, 1); +lean_inc(x_114); +lean_dec(x_103); +x_115 = lean_ctor_get(x_104, 1); +lean_inc(x_115); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_116 = x_104; +} else { + lean_dec_ref(x_104); + x_116 = lean_box(0); } +if (lean_is_scalar(x_116)) { + x_117 = lean_alloc_ctor(0, 2, 0); +} else { + x_117 = x_116; } +lean_ctor_set(x_117, 0, x_1); +lean_ctor_set(x_117, 1, x_115); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_114); +return x_118; } -block_42: -{ -lean_object* x_13; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_12); -x_13 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_ctor_get(x_17, 2); -lean_inc(x_18); -lean_inc(x_18); -x_19 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_18, x_16); -if (lean_obj_tag(x_19) == 0) -{ -size_t x_20; size_t x_21; uint8_t x_22; -x_20 = lean_ptr_addr(x_12); -lean_dec(x_12); -x_21 = lean_ptr_addr(x_16); -x_22 = lean_usize_dec_eq(x_20, x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_23 = lean_unsigned_to_nat(0u); -lean_inc(x_16); -x_24 = lean_array_set(x_3, x_23, x_16); -x_25 = l_Lean_mkAppN(x_2, x_24); -lean_dec(x_24); -x_26 = lean_ctor_get(x_17, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_17, 1); -lean_inc(x_27); -lean_dec(x_17); -lean_inc(x_25); -x_28 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_18, x_16, x_25); -x_29 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_29, 0, x_26); -lean_ctor_set(x_29, 1, x_27); -lean_ctor_set(x_29, 2, x_28); -x_30 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_25, x_5, x_29, x_7, x_8, x_9, x_10, x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_30; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_dec(x_3); -lean_dec(x_2); -x_31 = lean_ctor_get(x_17, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_17, 1); -lean_inc(x_32); -lean_dec(x_17); -lean_inc(x_1); -x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_18, x_16, x_1); -x_34 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_34, 0, x_31); -lean_ctor_set(x_34, 1, x_32); -lean_ctor_set(x_34, 2, x_33); -lean_inc(x_1); -x_35 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_34, x_7, x_8, x_9, x_10, x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_35; -} -} -else +uint8_t x_119; +lean_dec(x_1); +x_119 = !lean_is_exclusive(x_103); +if (x_119 == 0) { -lean_object* x_36; lean_object* x_37; -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_12); -lean_dec(x_3); -lean_dec(x_2); -x_36 = lean_ctor_get(x_19, 0); -lean_inc(x_36); -lean_dec(x_19); -x_37 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_36, x_5, x_17, x_7, x_8, x_9, x_10, x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_37; -} +lean_object* x_120; uint8_t x_121; +x_120 = lean_ctor_get(x_103, 0); +lean_dec(x_120); +x_121 = !lean_is_exclusive(x_104); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_104, 0); +lean_dec(x_122); +x_123 = lean_ctor_get(x_105, 0); +lean_inc(x_123); +lean_dec(x_105); +x_124 = l_Lean_mkAppN(x_2, x_123); +lean_dec(x_123); +lean_ctor_set(x_104, 0, x_124); +return x_103; } else { -uint8_t x_38; -lean_dec(x_12); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_38 = !lean_is_exclusive(x_13); -if (x_38 == 0) -{ -return x_13; +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_125 = lean_ctor_get(x_104, 1); +lean_inc(x_125); +lean_dec(x_104); +x_126 = lean_ctor_get(x_105, 0); +lean_inc(x_126); +lean_dec(x_105); +x_127 = l_Lean_mkAppN(x_2, x_126); +lean_dec(x_126); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_125); +lean_ctor_set(x_103, 0, x_128); +return x_103; +} } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_13, 0); -x_40 = lean_ctor_get(x_13, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_13); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; -} +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_129 = lean_ctor_get(x_103, 1); +lean_inc(x_129); +lean_dec(x_103); +x_130 = lean_ctor_get(x_104, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_131 = x_104; +} else { + lean_dec_ref(x_104); + x_131 = lean_box(0); } +x_132 = lean_ctor_get(x_105, 0); +lean_inc(x_132); +lean_dec(x_105); +x_133 = l_Lean_mkAppN(x_2, x_132); +lean_dec(x_132); +if (lean_is_scalar(x_131)) { + x_134 = lean_alloc_ctor(0, 2, 0); +} else { + x_134 = x_131; } -block_76: -{ -lean_object* x_44; -lean_dec(x_43); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_44 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_44) == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = lean_ctor_get(x_45, 0); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_array_get_size(x_3); -x_49 = lean_unsigned_to_nat(0u); -x_50 = lean_unsigned_to_nat(1u); -x_51 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_48); -lean_ctor_set(x_51, 2, x_50); -x_52 = 0; -x_53 = lean_box(x_52); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_3); -lean_ctor_set(x_54, 1, x_53); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_55 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_47, x_51, x_51, x_54, x_49, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_46); -lean_dec(x_51); -lean_dec(x_47); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -x_59 = lean_unbox(x_58); -lean_dec(x_58); -if (x_59 == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_57); -lean_dec(x_2); -x_60 = lean_ctor_get(x_55, 1); -lean_inc(x_60); -lean_dec(x_55); -x_61 = lean_ctor_get(x_56, 1); -lean_inc(x_61); -lean_dec(x_56); -lean_inc(x_1); -x_62 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_61, x_7, x_8, x_9, x_10, x_60); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_62; +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_130); +x_135 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_129); +return x_135; } -else -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_63 = lean_ctor_get(x_55, 1); -lean_inc(x_63); -lean_dec(x_55); -x_64 = lean_ctor_get(x_56, 1); -lean_inc(x_64); -lean_dec(x_56); -x_65 = lean_ctor_get(x_57, 0); -lean_inc(x_65); -lean_dec(x_57); -x_66 = l_Lean_mkAppN(x_2, x_65); -lean_dec(x_65); -x_67 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_66, x_5, x_64, x_7, x_8, x_9, x_10, x_63); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_67; } } else { -uint8_t x_68; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +uint8_t x_136; lean_dec(x_2); lean_dec(x_1); -x_68 = !lean_is_exclusive(x_55); -if (x_68 == 0) +x_136 = !lean_is_exclusive(x_103); +if (x_136 == 0) { -return x_55; +return x_103; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_55, 0); -x_70 = lean_ctor_get(x_55, 1); -lean_inc(x_70); -lean_inc(x_69); -lean_dec(x_55); -x_71 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -return x_71; +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_103, 0); +x_138 = lean_ctor_get(x_103, 1); +lean_inc(x_138); +lean_inc(x_137); +lean_dec(x_103); +x_139 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_139, 0, x_137); +lean_ctor_set(x_139, 1, x_138); +return x_139; } } } else { -uint8_t x_72; +uint8_t x_140; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7166,345 +6961,620 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = !lean_is_exclusive(x_44); -if (x_72 == 0) +x_140 = !lean_is_exclusive(x_92); +if (x_140 == 0) { -return x_44; +return x_92; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_44, 0); -x_74 = lean_ctor_get(x_44, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_44); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_92, 0); +x_142 = lean_ctor_get(x_92, 1); +lean_inc(x_142); +lean_inc(x_141); +lean_dec(x_92); +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +return x_143; } } } } case 1: { -lean_object* x_89; lean_object* x_120; lean_object* x_154; uint8_t x_155; +lean_object* x_157; lean_object* x_236; lean_object* x_290; uint8_t x_291; lean_dec(x_4); -x_154 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_155 = l_Lean_Expr_isConstOf(x_2, x_154); -if (x_155 == 0) +x_290 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_291 = l_Lean_Expr_isConstOf(x_2, x_290); +if (x_291 == 0) { -lean_object* x_156; -x_156 = lean_box(0); -x_120 = x_156; -goto block_153; +lean_object* x_292; +x_292 = lean_box(0); +x_236 = x_292; +goto block_289; } else { -lean_object* x_157; lean_object* x_158; uint8_t x_159; -x_157 = lean_array_get_size(x_3); -x_158 = lean_unsigned_to_nat(2u); -x_159 = lean_nat_dec_eq(x_157, x_158); -if (x_159 == 0) +lean_object* x_293; lean_object* x_294; uint8_t x_295; +x_293 = lean_array_get_size(x_3); +x_294 = lean_unsigned_to_nat(2u); +x_295 = lean_nat_dec_eq(x_293, x_294); +if (x_295 == 0) { -lean_object* x_160; -lean_dec(x_157); -x_160 = lean_box(0); -x_120 = x_160; -goto block_153; +lean_object* x_296; +lean_dec(x_293); +x_296 = lean_box(0); +x_236 = x_296; +goto block_289; } else { -lean_object* x_161; uint8_t x_162; -x_161 = lean_unsigned_to_nat(0u); -x_162 = lean_nat_dec_lt(x_161, x_157); -lean_dec(x_157); -if (x_162 == 0) +lean_object* x_297; uint8_t x_298; +x_297 = lean_unsigned_to_nat(0u); +x_298 = lean_nat_dec_lt(x_297, x_293); +lean_dec(x_293); +if (x_298 == 0) { -lean_object* x_163; lean_object* x_164; -x_163 = l_Lean_instInhabitedExpr; -x_164 = l_outOfBounds___rarg(x_163); -x_89 = x_164; -goto block_119; +lean_object* x_299; lean_object* x_300; +x_299 = l_Lean_instInhabitedExpr; +x_300 = l_outOfBounds___rarg(x_299); +x_157 = x_300; +goto block_235; } else { -lean_object* x_165; -x_165 = lean_array_fget(x_3, x_161); -x_89 = x_165; -goto block_119; +lean_object* x_301; +x_301 = lean_array_fget(x_3, x_297); +x_157 = x_301; +goto block_235; } } } -block_119: +block_235: { -lean_object* x_90; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_89); -x_90 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_89, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_90) == 0) +lean_object* x_158; +lean_inc(x_157); +x_158 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_157, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_158) == 0) { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_91 = lean_ctor_get(x_90, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_90, 1); -lean_inc(x_92); -lean_dec(x_90); -x_93 = lean_ctor_get(x_91, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_91, 1); -lean_inc(x_94); -lean_dec(x_91); -x_95 = lean_ctor_get(x_94, 2); -lean_inc(x_95); -lean_inc(x_95); -x_96 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_95, x_93); -if (lean_obj_tag(x_96) == 0) +uint8_t x_159; +x_159 = !lean_is_exclusive(x_158); +if (x_159 == 0) { -size_t x_97; size_t x_98; uint8_t x_99; -x_97 = lean_ptr_addr(x_89); -lean_dec(x_89); -x_98 = lean_ptr_addr(x_93); -x_99 = lean_usize_dec_eq(x_97, x_98); -if (x_99 == 0) +lean_object* x_160; uint8_t x_161; +x_160 = lean_ctor_get(x_158, 0); +x_161 = !lean_is_exclusive(x_160); +if (x_161 == 0) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_162 = lean_ctor_get(x_160, 0); +x_163 = lean_ctor_get(x_160, 1); +x_164 = lean_ctor_get(x_163, 2); +lean_inc(x_164); +lean_inc(x_164); +x_165 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_164, x_162); +if (lean_obj_tag(x_165) == 0) +{ +size_t x_166; size_t x_167; uint8_t x_168; +x_166 = lean_ptr_addr(x_157); +lean_dec(x_157); +x_167 = lean_ptr_addr(x_162); +x_168 = lean_usize_dec_eq(x_166, x_167); +if (x_168 == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_100 = lean_unsigned_to_nat(0u); -lean_inc(x_93); -x_101 = lean_array_set(x_3, x_100, x_93); -x_102 = l_Lean_mkAppN(x_2, x_101); -lean_dec(x_101); -x_103 = lean_ctor_get(x_94, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_94, 1); -lean_inc(x_104); -lean_dec(x_94); -lean_inc(x_102); -x_105 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_95, x_93, x_102); -x_106 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -lean_ctor_set(x_106, 2, x_105); -x_107 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_102, x_5, x_106, x_7, x_8, x_9, x_10, x_92); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_107; +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_1); +x_169 = lean_unsigned_to_nat(0u); +lean_inc(x_162); +x_170 = lean_array_set(x_3, x_169, x_162); +x_171 = l_Lean_mkAppN(x_2, x_170); +lean_dec(x_170); +x_172 = lean_ctor_get(x_163, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_163, 1); +lean_inc(x_173); +lean_dec(x_163); +lean_inc(x_171); +x_174 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_164, x_162, x_171); +x_175 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_175, 0, x_172); +lean_ctor_set(x_175, 1, x_173); +lean_ctor_set(x_175, 2, x_174); +lean_ctor_set(x_160, 1, x_175); +lean_ctor_set(x_160, 0, x_171); +return x_158; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_dec(x_3); lean_dec(x_2); -x_108 = lean_ctor_get(x_94, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_94, 1); -lean_inc(x_109); -lean_dec(x_94); -lean_inc(x_1); -x_110 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_95, x_93, x_1); -x_111 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -lean_ctor_set(x_111, 2, x_110); +x_176 = lean_ctor_get(x_163, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_163, 1); +lean_inc(x_177); +lean_dec(x_163); lean_inc(x_1); -x_112 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_111, x_7, x_8, x_9, x_10, x_92); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_112; +x_178 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_164, x_162, x_1); +x_179 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_179, 0, x_176); +lean_ctor_set(x_179, 1, x_177); +lean_ctor_set(x_179, 2, x_178); +lean_ctor_set(x_160, 1, x_179); +lean_ctor_set(x_160, 0, x_1); +return x_158; } } else { -lean_object* x_113; lean_object* x_114; -lean_dec(x_95); -lean_dec(x_93); -lean_dec(x_89); +lean_object* x_180; +lean_dec(x_164); +lean_dec(x_162); +lean_dec(x_157); lean_dec(x_3); lean_dec(x_2); -x_113 = lean_ctor_get(x_96, 0); -lean_inc(x_113); -lean_dec(x_96); -x_114 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_113, x_5, x_94, x_7, x_8, x_9, x_10, x_92); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_114; +lean_dec(x_1); +x_180 = lean_ctor_get(x_165, 0); +lean_inc(x_180); +lean_dec(x_165); +lean_ctor_set(x_160, 0, x_180); +return x_158; } } else { -uint8_t x_115; -lean_dec(x_89); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_181 = lean_ctor_get(x_160, 0); +x_182 = lean_ctor_get(x_160, 1); +lean_inc(x_182); +lean_inc(x_181); +lean_dec(x_160); +x_183 = lean_ctor_get(x_182, 2); +lean_inc(x_183); +lean_inc(x_183); +x_184 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_183, x_181); +if (lean_obj_tag(x_184) == 0) +{ +size_t x_185; size_t x_186; uint8_t x_187; +x_185 = lean_ptr_addr(x_157); +lean_dec(x_157); +x_186 = lean_ptr_addr(x_181); +x_187 = lean_usize_dec_eq(x_185, x_186); +if (x_187 == 0) +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +lean_dec(x_1); +x_188 = lean_unsigned_to_nat(0u); +lean_inc(x_181); +x_189 = lean_array_set(x_3, x_188, x_181); +x_190 = l_Lean_mkAppN(x_2, x_189); +lean_dec(x_189); +x_191 = lean_ctor_get(x_182, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_182, 1); +lean_inc(x_192); +lean_dec(x_182); +lean_inc(x_190); +x_193 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_183, x_181, x_190); +x_194 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_194, 0, x_191); +lean_ctor_set(x_194, 1, x_192); +lean_ctor_set(x_194, 2, x_193); +x_195 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_195, 0, x_190); +lean_ctor_set(x_195, 1, x_194); +lean_ctor_set(x_158, 0, x_195); +return x_158; +} +else +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_3); +lean_dec(x_2); +x_196 = lean_ctor_get(x_182, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_182, 1); +lean_inc(x_197); +lean_dec(x_182); +lean_inc(x_1); +x_198 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_183, x_181, x_1); +x_199 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_199, 0, x_196); +lean_ctor_set(x_199, 1, x_197); +lean_ctor_set(x_199, 2, x_198); +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_1); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_158, 0, x_200); +return x_158; +} +} +else +{ +lean_object* x_201; lean_object* x_202; +lean_dec(x_183); +lean_dec(x_181); +lean_dec(x_157); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_115 = !lean_is_exclusive(x_90); -if (x_115 == 0) +x_201 = lean_ctor_get(x_184, 0); +lean_inc(x_201); +lean_dec(x_184); +x_202 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_182); +lean_ctor_set(x_158, 0, x_202); +return x_158; +} +} +} +else +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_203 = lean_ctor_get(x_158, 0); +x_204 = lean_ctor_get(x_158, 1); +lean_inc(x_204); +lean_inc(x_203); +lean_dec(x_158); +x_205 = lean_ctor_get(x_203, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_203, 1); +lean_inc(x_206); +if (lean_is_exclusive(x_203)) { + lean_ctor_release(x_203, 0); + lean_ctor_release(x_203, 1); + x_207 = x_203; +} else { + lean_dec_ref(x_203); + x_207 = lean_box(0); +} +x_208 = lean_ctor_get(x_206, 2); +lean_inc(x_208); +lean_inc(x_208); +x_209 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_208, x_205); +if (lean_obj_tag(x_209) == 0) +{ +size_t x_210; size_t x_211; uint8_t x_212; +x_210 = lean_ptr_addr(x_157); +lean_dec(x_157); +x_211 = lean_ptr_addr(x_205); +x_212 = lean_usize_dec_eq(x_210, x_211); +if (x_212 == 0) { -return x_90; +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +lean_dec(x_1); +x_213 = lean_unsigned_to_nat(0u); +lean_inc(x_205); +x_214 = lean_array_set(x_3, x_213, x_205); +x_215 = l_Lean_mkAppN(x_2, x_214); +lean_dec(x_214); +x_216 = lean_ctor_get(x_206, 0); +lean_inc(x_216); +x_217 = lean_ctor_get(x_206, 1); +lean_inc(x_217); +lean_dec(x_206); +lean_inc(x_215); +x_218 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_208, x_205, x_215); +x_219 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_219, 0, x_216); +lean_ctor_set(x_219, 1, x_217); +lean_ctor_set(x_219, 2, x_218); +if (lean_is_scalar(x_207)) { + x_220 = lean_alloc_ctor(0, 2, 0); +} else { + x_220 = x_207; +} +lean_ctor_set(x_220, 0, x_215); +lean_ctor_set(x_220, 1, x_219); +x_221 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_221, 0, x_220); +lean_ctor_set(x_221, 1, x_204); +return x_221; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_116 = lean_ctor_get(x_90, 0); -x_117 = lean_ctor_get(x_90, 1); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_90); -x_118 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_118, 0, x_116); -lean_ctor_set(x_118, 1, x_117); -return x_118; +lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; +lean_dec(x_3); +lean_dec(x_2); +x_222 = lean_ctor_get(x_206, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_206, 1); +lean_inc(x_223); +lean_dec(x_206); +lean_inc(x_1); +x_224 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_208, x_205, x_1); +x_225 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_225, 0, x_222); +lean_ctor_set(x_225, 1, x_223); +lean_ctor_set(x_225, 2, x_224); +if (lean_is_scalar(x_207)) { + x_226 = lean_alloc_ctor(0, 2, 0); +} else { + x_226 = x_207; } +lean_ctor_set(x_226, 0, x_1); +lean_ctor_set(x_226, 1, x_225); +x_227 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_227, 0, x_226); +lean_ctor_set(x_227, 1, x_204); +return x_227; } } -block_153: +else { -lean_object* x_121; -lean_dec(x_120); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_121 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_121) == 0) +lean_object* x_228; lean_object* x_229; lean_object* x_230; +lean_dec(x_208); +lean_dec(x_205); +lean_dec(x_157); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_228 = lean_ctor_get(x_209, 0); +lean_inc(x_228); +lean_dec(x_209); +if (lean_is_scalar(x_207)) { + x_229 = lean_alloc_ctor(0, 2, 0); +} else { + x_229 = x_207; +} +lean_ctor_set(x_229, 0, x_228); +lean_ctor_set(x_229, 1, x_206); +x_230 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_230, 0, x_229); +lean_ctor_set(x_230, 1, x_204); +return x_230; +} +} +} +else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -x_123 = lean_ctor_get(x_121, 1); -lean_inc(x_123); -lean_dec(x_121); -x_124 = lean_ctor_get(x_122, 0); -lean_inc(x_124); -lean_dec(x_122); -x_125 = lean_array_get_size(x_3); -x_126 = lean_unsigned_to_nat(0u); -x_127 = lean_unsigned_to_nat(1u); -x_128 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_125); -lean_ctor_set(x_128, 2, x_127); -x_129 = 0; -x_130 = lean_box(x_129); -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_3); -lean_ctor_set(x_131, 1, x_130); +uint8_t x_231; +lean_dec(x_157); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_231 = !lean_is_exclusive(x_158); +if (x_231 == 0) +{ +return x_158; +} +else +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_232 = lean_ctor_get(x_158, 0); +x_233 = lean_ctor_get(x_158, 1); +lean_inc(x_233); +lean_inc(x_232); +lean_dec(x_158); +x_234 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_234, 0, x_232); +lean_ctor_set(x_234, 1, x_233); +return x_234; +} +} +} +block_289: +{ +lean_object* x_237; +lean_dec(x_236); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_5); lean_inc(x_2); -x_132 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_124, x_128, x_128, x_131, x_126, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_123); -lean_dec(x_128); -lean_dec(x_124); -if (lean_obj_tag(x_132) == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_134, 1); -lean_inc(x_135); -x_136 = lean_unbox(x_135); -lean_dec(x_135); -if (x_136 == 0) +x_237 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_237) == 0) +{ +lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; uint8_t x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_238 = lean_ctor_get(x_237, 0); +lean_inc(x_238); +x_239 = lean_ctor_get(x_237, 1); +lean_inc(x_239); +lean_dec(x_237); +x_240 = lean_ctor_get(x_238, 0); +lean_inc(x_240); +lean_dec(x_238); +x_241 = lean_array_get_size(x_3); +x_242 = lean_unsigned_to_nat(0u); +x_243 = lean_unsigned_to_nat(1u); +lean_inc(x_241); +x_244 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_244, 0, x_242); +lean_ctor_set(x_244, 1, x_241); +lean_ctor_set(x_244, 2, x_243); +x_245 = 0; +x_246 = lean_box(x_245); +lean_inc(x_3); +x_247 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_247, 0, x_3); +lean_ctor_set(x_247, 1, x_246); +lean_inc(x_2); +x_248 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_240, x_241, x_244, x_244, x_247, x_242, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_239); +lean_dec(x_244); +lean_dec(x_241); +lean_dec(x_240); +lean_dec(x_3); +if (lean_obj_tag(x_248) == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -lean_dec(x_134); +lean_object* x_249; lean_object* x_250; lean_object* x_251; uint8_t x_252; +x_249 = lean_ctor_get(x_248, 0); +lean_inc(x_249); +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_250, 1); +lean_inc(x_251); +x_252 = lean_unbox(x_251); +lean_dec(x_251); +if (x_252 == 0) +{ +uint8_t x_253; +lean_dec(x_250); lean_dec(x_2); -x_137 = lean_ctor_get(x_132, 1); -lean_inc(x_137); -lean_dec(x_132); -x_138 = lean_ctor_get(x_133, 1); -lean_inc(x_138); -lean_dec(x_133); -lean_inc(x_1); -x_139 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_138, x_7, x_8, x_9, x_10, x_137); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_139; +x_253 = !lean_is_exclusive(x_248); +if (x_253 == 0) +{ +lean_object* x_254; uint8_t x_255; +x_254 = lean_ctor_get(x_248, 0); +lean_dec(x_254); +x_255 = !lean_is_exclusive(x_249); +if (x_255 == 0) +{ +lean_object* x_256; +x_256 = lean_ctor_get(x_249, 0); +lean_dec(x_256); +lean_ctor_set(x_249, 0, x_1); +return x_248; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_140 = lean_ctor_get(x_132, 1); -lean_inc(x_140); -lean_dec(x_132); -x_141 = lean_ctor_get(x_133, 1); -lean_inc(x_141); -lean_dec(x_133); -x_142 = lean_ctor_get(x_134, 0); -lean_inc(x_142); -lean_dec(x_134); -x_143 = l_Lean_mkAppN(x_2, x_142); -lean_dec(x_142); -x_144 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_143, x_5, x_141, x_7, x_8, x_9, x_10, x_140); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_144; +lean_object* x_257; lean_object* x_258; +x_257 = lean_ctor_get(x_249, 1); +lean_inc(x_257); +lean_dec(x_249); +x_258 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_258, 0, x_1); +lean_ctor_set(x_258, 1, x_257); +lean_ctor_set(x_248, 0, x_258); +return x_248; } } else { -uint8_t x_145; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_259 = lean_ctor_get(x_248, 1); +lean_inc(x_259); +lean_dec(x_248); +x_260 = lean_ctor_get(x_249, 1); +lean_inc(x_260); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_261 = x_249; +} else { + lean_dec_ref(x_249); + x_261 = lean_box(0); +} +if (lean_is_scalar(x_261)) { + x_262 = lean_alloc_ctor(0, 2, 0); +} else { + x_262 = x_261; +} +lean_ctor_set(x_262, 0, x_1); +lean_ctor_set(x_262, 1, x_260); +x_263 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_263, 0, x_262); +lean_ctor_set(x_263, 1, x_259); +return x_263; +} +} +else +{ +uint8_t x_264; +lean_dec(x_1); +x_264 = !lean_is_exclusive(x_248); +if (x_264 == 0) +{ +lean_object* x_265; uint8_t x_266; +x_265 = lean_ctor_get(x_248, 0); +lean_dec(x_265); +x_266 = !lean_is_exclusive(x_249); +if (x_266 == 0) +{ +lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_267 = lean_ctor_get(x_249, 0); +lean_dec(x_267); +x_268 = lean_ctor_get(x_250, 0); +lean_inc(x_268); +lean_dec(x_250); +x_269 = l_Lean_mkAppN(x_2, x_268); +lean_dec(x_268); +lean_ctor_set(x_249, 0, x_269); +return x_248; +} +else +{ +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +x_270 = lean_ctor_get(x_249, 1); +lean_inc(x_270); +lean_dec(x_249); +x_271 = lean_ctor_get(x_250, 0); +lean_inc(x_271); +lean_dec(x_250); +x_272 = l_Lean_mkAppN(x_2, x_271); +lean_dec(x_271); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_272); +lean_ctor_set(x_273, 1, x_270); +lean_ctor_set(x_248, 0, x_273); +return x_248; +} +} +else +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; +x_274 = lean_ctor_get(x_248, 1); +lean_inc(x_274); +lean_dec(x_248); +x_275 = lean_ctor_get(x_249, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_276 = x_249; +} else { + lean_dec_ref(x_249); + x_276 = lean_box(0); +} +x_277 = lean_ctor_get(x_250, 0); +lean_inc(x_277); +lean_dec(x_250); +x_278 = l_Lean_mkAppN(x_2, x_277); +lean_dec(x_277); +if (lean_is_scalar(x_276)) { + x_279 = lean_alloc_ctor(0, 2, 0); +} else { + x_279 = x_276; +} +lean_ctor_set(x_279, 0, x_278); +lean_ctor_set(x_279, 1, x_275); +x_280 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_280, 0, x_279); +lean_ctor_set(x_280, 1, x_274); +return x_280; +} +} +} +else +{ +uint8_t x_281; lean_dec(x_2); lean_dec(x_1); -x_145 = !lean_is_exclusive(x_132); -if (x_145 == 0) +x_281 = !lean_is_exclusive(x_248); +if (x_281 == 0) { -return x_132; +return x_248; } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_132, 0); -x_147 = lean_ctor_get(x_132, 1); -lean_inc(x_147); -lean_inc(x_146); -lean_dec(x_132); -x_148 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_147); -return x_148; +lean_object* x_282; lean_object* x_283; lean_object* x_284; +x_282 = lean_ctor_get(x_248, 0); +x_283 = lean_ctor_get(x_248, 1); +lean_inc(x_283); +lean_inc(x_282); +lean_dec(x_248); +x_284 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_284, 0, x_282); +lean_ctor_set(x_284, 1, x_283); +return x_284; } } } else { -uint8_t x_149; +uint8_t x_285; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7514,693 +7584,620 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_149 = !lean_is_exclusive(x_121); -if (x_149 == 0) +x_285 = !lean_is_exclusive(x_237); +if (x_285 == 0) { -return x_121; +return x_237; } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_150 = lean_ctor_get(x_121, 0); -x_151 = lean_ctor_get(x_121, 1); -lean_inc(x_151); -lean_inc(x_150); -lean_dec(x_121); -x_152 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_152, 0, x_150); -lean_ctor_set(x_152, 1, x_151); -return x_152; +lean_object* x_286; lean_object* x_287; lean_object* x_288; +x_286 = lean_ctor_get(x_237, 0); +x_287 = lean_ctor_get(x_237, 1); +lean_inc(x_287); +lean_inc(x_286); +lean_dec(x_237); +x_288 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_288, 0, x_286); +lean_ctor_set(x_288, 1, x_287); +return x_288; } } } } case 2: { -lean_object* x_166; lean_object* x_197; lean_object* x_231; uint8_t x_232; +lean_object* x_302; lean_object* x_381; lean_object* x_435; uint8_t x_436; lean_dec(x_4); -x_231 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_232 = l_Lean_Expr_isConstOf(x_2, x_231); -if (x_232 == 0) +x_435 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_436 = l_Lean_Expr_isConstOf(x_2, x_435); +if (x_436 == 0) { -lean_object* x_233; -x_233 = lean_box(0); -x_197 = x_233; -goto block_230; +lean_object* x_437; +x_437 = lean_box(0); +x_381 = x_437; +goto block_434; } else { -lean_object* x_234; lean_object* x_235; uint8_t x_236; -x_234 = lean_array_get_size(x_3); -x_235 = lean_unsigned_to_nat(2u); -x_236 = lean_nat_dec_eq(x_234, x_235); -if (x_236 == 0) +lean_object* x_438; lean_object* x_439; uint8_t x_440; +x_438 = lean_array_get_size(x_3); +x_439 = lean_unsigned_to_nat(2u); +x_440 = lean_nat_dec_eq(x_438, x_439); +if (x_440 == 0) { -lean_object* x_237; -lean_dec(x_234); -x_237 = lean_box(0); -x_197 = x_237; -goto block_230; +lean_object* x_441; +lean_dec(x_438); +x_441 = lean_box(0); +x_381 = x_441; +goto block_434; } else { -lean_object* x_238; uint8_t x_239; -x_238 = lean_unsigned_to_nat(0u); -x_239 = lean_nat_dec_lt(x_238, x_234); -lean_dec(x_234); -if (x_239 == 0) +lean_object* x_442; uint8_t x_443; +x_442 = lean_unsigned_to_nat(0u); +x_443 = lean_nat_dec_lt(x_442, x_438); +lean_dec(x_438); +if (x_443 == 0) { -lean_object* x_240; lean_object* x_241; -x_240 = l_Lean_instInhabitedExpr; -x_241 = l_outOfBounds___rarg(x_240); -x_166 = x_241; -goto block_196; +lean_object* x_444; lean_object* x_445; +x_444 = l_Lean_instInhabitedExpr; +x_445 = l_outOfBounds___rarg(x_444); +x_302 = x_445; +goto block_380; } else { -lean_object* x_242; -x_242 = lean_array_fget(x_3, x_238); -x_166 = x_242; -goto block_196; +lean_object* x_446; +x_446 = lean_array_fget(x_3, x_442); +x_302 = x_446; +goto block_380; } } } -block_196: +block_380: { -lean_object* x_167; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_166); -x_167 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_166, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_167) == 0) +lean_object* x_303; +lean_inc(x_302); +x_303 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_302, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_303) == 0) { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_168 = lean_ctor_get(x_167, 0); -lean_inc(x_168); -x_169 = lean_ctor_get(x_167, 1); -lean_inc(x_169); -lean_dec(x_167); -x_170 = lean_ctor_get(x_168, 0); -lean_inc(x_170); -x_171 = lean_ctor_get(x_168, 1); -lean_inc(x_171); -lean_dec(x_168); -x_172 = lean_ctor_get(x_171, 2); -lean_inc(x_172); -lean_inc(x_172); -x_173 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_172, x_170); -if (lean_obj_tag(x_173) == 0) -{ -size_t x_174; size_t x_175; uint8_t x_176; -x_174 = lean_ptr_addr(x_166); -lean_dec(x_166); -x_175 = lean_ptr_addr(x_170); -x_176 = lean_usize_dec_eq(x_174, x_175); -if (x_176 == 0) -{ -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_177 = lean_unsigned_to_nat(0u); -lean_inc(x_170); -x_178 = lean_array_set(x_3, x_177, x_170); -x_179 = l_Lean_mkAppN(x_2, x_178); -lean_dec(x_178); -x_180 = lean_ctor_get(x_171, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_171, 1); -lean_inc(x_181); -lean_dec(x_171); -lean_inc(x_179); -x_182 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_172, x_170, x_179); -x_183 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_183, 0, x_180); -lean_ctor_set(x_183, 1, x_181); -lean_ctor_set(x_183, 2, x_182); -x_184 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_179, x_5, x_183, x_7, x_8, x_9, x_10, x_169); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_184; -} -else +uint8_t x_304; +x_304 = !lean_is_exclusive(x_303); +if (x_304 == 0) +{ +lean_object* x_305; uint8_t x_306; +x_305 = lean_ctor_get(x_303, 0); +x_306 = !lean_is_exclusive(x_305); +if (x_306 == 0) +{ +lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +x_307 = lean_ctor_get(x_305, 0); +x_308 = lean_ctor_get(x_305, 1); +x_309 = lean_ctor_get(x_308, 2); +lean_inc(x_309); +lean_inc(x_309); +x_310 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_309, x_307); +if (lean_obj_tag(x_310) == 0) +{ +size_t x_311; size_t x_312; uint8_t x_313; +x_311 = lean_ptr_addr(x_302); +lean_dec(x_302); +x_312 = lean_ptr_addr(x_307); +x_313 = lean_usize_dec_eq(x_311, x_312); +if (x_313 == 0) { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; +lean_dec(x_1); +x_314 = lean_unsigned_to_nat(0u); +lean_inc(x_307); +x_315 = lean_array_set(x_3, x_314, x_307); +x_316 = l_Lean_mkAppN(x_2, x_315); +lean_dec(x_315); +x_317 = lean_ctor_get(x_308, 0); +lean_inc(x_317); +x_318 = lean_ctor_get(x_308, 1); +lean_inc(x_318); +lean_dec(x_308); +lean_inc(x_316); +x_319 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_309, x_307, x_316); +x_320 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_320, 0, x_317); +lean_ctor_set(x_320, 1, x_318); +lean_ctor_set(x_320, 2, x_319); +lean_ctor_set(x_305, 1, x_320); +lean_ctor_set(x_305, 0, x_316); +return x_303; +} +else +{ +lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_dec(x_3); lean_dec(x_2); -x_185 = lean_ctor_get(x_171, 0); -lean_inc(x_185); -x_186 = lean_ctor_get(x_171, 1); -lean_inc(x_186); -lean_dec(x_171); -lean_inc(x_1); -x_187 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_172, x_170, x_1); -x_188 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_188, 0, x_185); -lean_ctor_set(x_188, 1, x_186); -lean_ctor_set(x_188, 2, x_187); +x_321 = lean_ctor_get(x_308, 0); +lean_inc(x_321); +x_322 = lean_ctor_get(x_308, 1); +lean_inc(x_322); +lean_dec(x_308); lean_inc(x_1); -x_189 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_188, x_7, x_8, x_9, x_10, x_169); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_189; +x_323 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_309, x_307, x_1); +x_324 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_324, 0, x_321); +lean_ctor_set(x_324, 1, x_322); +lean_ctor_set(x_324, 2, x_323); +lean_ctor_set(x_305, 1, x_324); +lean_ctor_set(x_305, 0, x_1); +return x_303; } } else { -lean_object* x_190; lean_object* x_191; -lean_dec(x_172); -lean_dec(x_170); -lean_dec(x_166); +lean_object* x_325; +lean_dec(x_309); +lean_dec(x_307); +lean_dec(x_302); lean_dec(x_3); lean_dec(x_2); -x_190 = lean_ctor_get(x_173, 0); -lean_inc(x_190); -lean_dec(x_173); -x_191 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_190, x_5, x_171, x_7, x_8, x_9, x_10, x_169); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_191; +lean_dec(x_1); +x_325 = lean_ctor_get(x_310, 0); +lean_inc(x_325); +lean_dec(x_310); +lean_ctor_set(x_305, 0, x_325); +return x_303; } } else { -uint8_t x_192; -lean_dec(x_166); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; +x_326 = lean_ctor_get(x_305, 0); +x_327 = lean_ctor_get(x_305, 1); +lean_inc(x_327); +lean_inc(x_326); +lean_dec(x_305); +x_328 = lean_ctor_get(x_327, 2); +lean_inc(x_328); +lean_inc(x_328); +x_329 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_328, x_326); +if (lean_obj_tag(x_329) == 0) +{ +size_t x_330; size_t x_331; uint8_t x_332; +x_330 = lean_ptr_addr(x_302); +lean_dec(x_302); +x_331 = lean_ptr_addr(x_326); +x_332 = lean_usize_dec_eq(x_330, x_331); +if (x_332 == 0) +{ +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_dec(x_1); -x_192 = !lean_is_exclusive(x_167); -if (x_192 == 0) +x_333 = lean_unsigned_to_nat(0u); +lean_inc(x_326); +x_334 = lean_array_set(x_3, x_333, x_326); +x_335 = l_Lean_mkAppN(x_2, x_334); +lean_dec(x_334); +x_336 = lean_ctor_get(x_327, 0); +lean_inc(x_336); +x_337 = lean_ctor_get(x_327, 1); +lean_inc(x_337); +lean_dec(x_327); +lean_inc(x_335); +x_338 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_328, x_326, x_335); +x_339 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_339, 0, x_336); +lean_ctor_set(x_339, 1, x_337); +lean_ctor_set(x_339, 2, x_338); +x_340 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_340, 0, x_335); +lean_ctor_set(x_340, 1, x_339); +lean_ctor_set(x_303, 0, x_340); +return x_303; +} +else { -return x_167; +lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_3); +lean_dec(x_2); +x_341 = lean_ctor_get(x_327, 0); +lean_inc(x_341); +x_342 = lean_ctor_get(x_327, 1); +lean_inc(x_342); +lean_dec(x_327); +lean_inc(x_1); +x_343 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_328, x_326, x_1); +x_344 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_344, 0, x_341); +lean_ctor_set(x_344, 1, x_342); +lean_ctor_set(x_344, 2, x_343); +x_345 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_345, 0, x_1); +lean_ctor_set(x_345, 1, x_344); +lean_ctor_set(x_303, 0, x_345); +return x_303; +} } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_193 = lean_ctor_get(x_167, 0); -x_194 = lean_ctor_get(x_167, 1); -lean_inc(x_194); -lean_inc(x_193); -lean_dec(x_167); -x_195 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -return x_195; +lean_object* x_346; lean_object* x_347; +lean_dec(x_328); +lean_dec(x_326); +lean_dec(x_302); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_346 = lean_ctor_get(x_329, 0); +lean_inc(x_346); +lean_dec(x_329); +x_347 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_347, 0, x_346); +lean_ctor_set(x_347, 1, x_327); +lean_ctor_set(x_303, 0, x_347); +return x_303; } } } -block_230: +else { -lean_object* x_198; -lean_dec(x_197); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_198 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_198) == 0) -{ -lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; uint8_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -x_199 = lean_ctor_get(x_198, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_198, 1); -lean_inc(x_200); -lean_dec(x_198); -x_201 = lean_ctor_get(x_199, 0); -lean_inc(x_201); -lean_dec(x_199); -x_202 = lean_array_get_size(x_3); -x_203 = lean_unsigned_to_nat(0u); -x_204 = lean_unsigned_to_nat(1u); -x_205 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_202); -lean_ctor_set(x_205, 2, x_204); -x_206 = 0; -x_207 = lean_box(x_206); -x_208 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_208, 0, x_3); -lean_ctor_set(x_208, 1, x_207); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_209 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_201, x_205, x_205, x_208, x_203, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_200); -lean_dec(x_205); -lean_dec(x_201); -if (lean_obj_tag(x_209) == 0) -{ -lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; -x_210 = lean_ctor_get(x_209, 0); -lean_inc(x_210); -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_211, 1); -lean_inc(x_212); -x_213 = lean_unbox(x_212); -lean_dec(x_212); -if (x_213 == 0) -{ -lean_object* x_214; lean_object* x_215; lean_object* x_216; -lean_dec(x_211); -lean_dec(x_2); -x_214 = lean_ctor_get(x_209, 1); -lean_inc(x_214); -lean_dec(x_209); -x_215 = lean_ctor_get(x_210, 1); -lean_inc(x_215); -lean_dec(x_210); -lean_inc(x_1); -x_216 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_215, x_7, x_8, x_9, x_10, x_214); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_216; +lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; +x_348 = lean_ctor_get(x_303, 0); +x_349 = lean_ctor_get(x_303, 1); +lean_inc(x_349); +lean_inc(x_348); +lean_dec(x_303); +x_350 = lean_ctor_get(x_348, 0); +lean_inc(x_350); +x_351 = lean_ctor_get(x_348, 1); +lean_inc(x_351); +if (lean_is_exclusive(x_348)) { + lean_ctor_release(x_348, 0); + lean_ctor_release(x_348, 1); + x_352 = x_348; +} else { + lean_dec_ref(x_348); + x_352 = lean_box(0); } -else +x_353 = lean_ctor_get(x_351, 2); +lean_inc(x_353); +lean_inc(x_353); +x_354 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_353, x_350); +if (lean_obj_tag(x_354) == 0) { -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_217 = lean_ctor_get(x_209, 1); -lean_inc(x_217); -lean_dec(x_209); -x_218 = lean_ctor_get(x_210, 1); -lean_inc(x_218); -lean_dec(x_210); -x_219 = lean_ctor_get(x_211, 0); -lean_inc(x_219); -lean_dec(x_211); -x_220 = l_Lean_mkAppN(x_2, x_219); -lean_dec(x_219); -x_221 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_220, x_5, x_218, x_7, x_8, x_9, x_10, x_217); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_221; -} -} -else +size_t x_355; size_t x_356; uint8_t x_357; +x_355 = lean_ptr_addr(x_302); +lean_dec(x_302); +x_356 = lean_ptr_addr(x_350); +x_357 = lean_usize_dec_eq(x_355, x_356); +if (x_357 == 0) { -uint8_t x_222; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); +lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_dec(x_1); -x_222 = !lean_is_exclusive(x_209); -if (x_222 == 0) -{ -return x_209; +x_358 = lean_unsigned_to_nat(0u); +lean_inc(x_350); +x_359 = lean_array_set(x_3, x_358, x_350); +x_360 = l_Lean_mkAppN(x_2, x_359); +lean_dec(x_359); +x_361 = lean_ctor_get(x_351, 0); +lean_inc(x_361); +x_362 = lean_ctor_get(x_351, 1); +lean_inc(x_362); +lean_dec(x_351); +lean_inc(x_360); +x_363 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_353, x_350, x_360); +x_364 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_364, 0, x_361); +lean_ctor_set(x_364, 1, x_362); +lean_ctor_set(x_364, 2, x_363); +if (lean_is_scalar(x_352)) { + x_365 = lean_alloc_ctor(0, 2, 0); +} else { + x_365 = x_352; +} +lean_ctor_set(x_365, 0, x_360); +lean_ctor_set(x_365, 1, x_364); +x_366 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_366, 0, x_365); +lean_ctor_set(x_366, 1, x_349); +return x_366; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_223 = lean_ctor_get(x_209, 0); -x_224 = lean_ctor_get(x_209, 1); -lean_inc(x_224); -lean_inc(x_223); -lean_dec(x_209); -x_225 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -return x_225; +lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +lean_dec(x_3); +lean_dec(x_2); +x_367 = lean_ctor_get(x_351, 0); +lean_inc(x_367); +x_368 = lean_ctor_get(x_351, 1); +lean_inc(x_368); +lean_dec(x_351); +lean_inc(x_1); +x_369 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_353, x_350, x_1); +x_370 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_370, 0, x_367); +lean_ctor_set(x_370, 1, x_368); +lean_ctor_set(x_370, 2, x_369); +if (lean_is_scalar(x_352)) { + x_371 = lean_alloc_ctor(0, 2, 0); +} else { + x_371 = x_352; } +lean_ctor_set(x_371, 0, x_1); +lean_ctor_set(x_371, 1, x_370); +x_372 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_372, 0, x_371); +lean_ctor_set(x_372, 1, x_349); +return x_372; } } else { -uint8_t x_226; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_373; lean_object* x_374; lean_object* x_375; +lean_dec(x_353); +lean_dec(x_350); +lean_dec(x_302); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_226 = !lean_is_exclusive(x_198); -if (x_226 == 0) -{ -return x_198; -} -else -{ -lean_object* x_227; lean_object* x_228; lean_object* x_229; -x_227 = lean_ctor_get(x_198, 0); -x_228 = lean_ctor_get(x_198, 1); -lean_inc(x_228); -lean_inc(x_227); -lean_dec(x_198); -x_229 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_229, 0, x_227); -lean_ctor_set(x_229, 1, x_228); -return x_229; -} -} +x_373 = lean_ctor_get(x_354, 0); +lean_inc(x_373); +lean_dec(x_354); +if (lean_is_scalar(x_352)) { + x_374 = lean_alloc_ctor(0, 2, 0); +} else { + x_374 = x_352; } +lean_ctor_set(x_374, 0, x_373); +lean_ctor_set(x_374, 1, x_351); +x_375 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_375, 0, x_374); +lean_ctor_set(x_375, 1, x_349); +return x_375; } -case 3: -{ -lean_object* x_243; lean_object* x_274; lean_object* x_308; uint8_t x_309; -lean_dec(x_4); -x_308 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_309 = l_Lean_Expr_isConstOf(x_2, x_308); -if (x_309 == 0) -{ -lean_object* x_310; -x_310 = lean_box(0); -x_274 = x_310; -goto block_307; } -else -{ -lean_object* x_311; lean_object* x_312; uint8_t x_313; -x_311 = lean_array_get_size(x_3); -x_312 = lean_unsigned_to_nat(2u); -x_313 = lean_nat_dec_eq(x_311, x_312); -if (x_313 == 0) -{ -lean_object* x_314; -lean_dec(x_311); -x_314 = lean_box(0); -x_274 = x_314; -goto block_307; } else { -lean_object* x_315; uint8_t x_316; -x_315 = lean_unsigned_to_nat(0u); -x_316 = lean_nat_dec_lt(x_315, x_311); -lean_dec(x_311); -if (x_316 == 0) +uint8_t x_376; +lean_dec(x_302); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_376 = !lean_is_exclusive(x_303); +if (x_376 == 0) { -lean_object* x_317; lean_object* x_318; -x_317 = l_Lean_instInhabitedExpr; -x_318 = l_outOfBounds___rarg(x_317); -x_243 = x_318; -goto block_273; +return x_303; } else { -lean_object* x_319; -x_319 = lean_array_fget(x_3, x_315); -x_243 = x_319; -goto block_273; +lean_object* x_377; lean_object* x_378; lean_object* x_379; +x_377 = lean_ctor_get(x_303, 0); +x_378 = lean_ctor_get(x_303, 1); +lean_inc(x_378); +lean_inc(x_377); +lean_dec(x_303); +x_379 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_379, 0, x_377); +lean_ctor_set(x_379, 1, x_378); +return x_379; } } } -block_273: +block_434: { -lean_object* x_244; +lean_object* x_382; +lean_dec(x_381); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_243); -x_244 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_243, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_244) == 0) -{ -lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_245 = lean_ctor_get(x_244, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_244, 1); -lean_inc(x_246); -lean_dec(x_244); -x_247 = lean_ctor_get(x_245, 0); -lean_inc(x_247); -x_248 = lean_ctor_get(x_245, 1); -lean_inc(x_248); -lean_dec(x_245); -x_249 = lean_ctor_get(x_248, 2); -lean_inc(x_249); -lean_inc(x_249); -x_250 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_249, x_247); -if (lean_obj_tag(x_250) == 0) -{ -size_t x_251; size_t x_252; uint8_t x_253; -x_251 = lean_ptr_addr(x_243); -lean_dec(x_243); -x_252 = lean_ptr_addr(x_247); -x_253 = lean_usize_dec_eq(x_251, x_252); -if (x_253 == 0) +lean_inc(x_2); +x_382 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_382) == 0) +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; uint8_t x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; +x_383 = lean_ctor_get(x_382, 0); +lean_inc(x_383); +x_384 = lean_ctor_get(x_382, 1); +lean_inc(x_384); +lean_dec(x_382); +x_385 = lean_ctor_get(x_383, 0); +lean_inc(x_385); +lean_dec(x_383); +x_386 = lean_array_get_size(x_3); +x_387 = lean_unsigned_to_nat(0u); +x_388 = lean_unsigned_to_nat(1u); +lean_inc(x_386); +x_389 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_389, 0, x_387); +lean_ctor_set(x_389, 1, x_386); +lean_ctor_set(x_389, 2, x_388); +x_390 = 0; +x_391 = lean_box(x_390); +lean_inc(x_3); +x_392 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_392, 0, x_3); +lean_ctor_set(x_392, 1, x_391); +lean_inc(x_2); +x_393 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_385, x_386, x_389, x_389, x_392, x_387, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_384); +lean_dec(x_389); +lean_dec(x_386); +lean_dec(x_385); +lean_dec(x_3); +if (lean_obj_tag(x_393) == 0) +{ +lean_object* x_394; lean_object* x_395; lean_object* x_396; uint8_t x_397; +x_394 = lean_ctor_get(x_393, 0); +lean_inc(x_394); +x_395 = lean_ctor_get(x_394, 0); +lean_inc(x_395); +x_396 = lean_ctor_get(x_395, 1); +lean_inc(x_396); +x_397 = lean_unbox(x_396); +lean_dec(x_396); +if (x_397 == 0) +{ +uint8_t x_398; +lean_dec(x_395); +lean_dec(x_2); +x_398 = !lean_is_exclusive(x_393); +if (x_398 == 0) { -lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_254 = lean_unsigned_to_nat(0u); -lean_inc(x_247); -x_255 = lean_array_set(x_3, x_254, x_247); -x_256 = l_Lean_mkAppN(x_2, x_255); -lean_dec(x_255); -x_257 = lean_ctor_get(x_248, 0); -lean_inc(x_257); -x_258 = lean_ctor_get(x_248, 1); -lean_inc(x_258); -lean_dec(x_248); -lean_inc(x_256); -x_259 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_249, x_247, x_256); -x_260 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_260, 0, x_257); -lean_ctor_set(x_260, 1, x_258); -lean_ctor_set(x_260, 2, x_259); -x_261 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_256, x_5, x_260, x_7, x_8, x_9, x_10, x_246); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_261; +lean_object* x_399; uint8_t x_400; +x_399 = lean_ctor_get(x_393, 0); +lean_dec(x_399); +x_400 = !lean_is_exclusive(x_394); +if (x_400 == 0) +{ +lean_object* x_401; +x_401 = lean_ctor_get(x_394, 0); +lean_dec(x_401); +lean_ctor_set(x_394, 0, x_1); +return x_393; } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; -lean_dec(x_3); -lean_dec(x_2); -x_262 = lean_ctor_get(x_248, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_248, 1); -lean_inc(x_263); -lean_dec(x_248); -lean_inc(x_1); -x_264 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_249, x_247, x_1); -x_265 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_265, 0, x_262); -lean_ctor_set(x_265, 1, x_263); -lean_ctor_set(x_265, 2, x_264); -lean_inc(x_1); -x_266 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_265, x_7, x_8, x_9, x_10, x_246); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_266; +lean_object* x_402; lean_object* x_403; +x_402 = lean_ctor_get(x_394, 1); +lean_inc(x_402); +lean_dec(x_394); +x_403 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_403, 0, x_1); +lean_ctor_set(x_403, 1, x_402); +lean_ctor_set(x_393, 0, x_403); +return x_393; } } else { -lean_object* x_267; lean_object* x_268; -lean_dec(x_249); -lean_dec(x_247); -lean_dec(x_243); -lean_dec(x_3); -lean_dec(x_2); -x_267 = lean_ctor_get(x_250, 0); -lean_inc(x_267); -lean_dec(x_250); -x_268 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_267, x_5, x_248, x_7, x_8, x_9, x_10, x_246); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_268; +lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; +x_404 = lean_ctor_get(x_393, 1); +lean_inc(x_404); +lean_dec(x_393); +x_405 = lean_ctor_get(x_394, 1); +lean_inc(x_405); +if (lean_is_exclusive(x_394)) { + lean_ctor_release(x_394, 0); + lean_ctor_release(x_394, 1); + x_406 = x_394; +} else { + lean_dec_ref(x_394); + x_406 = lean_box(0); +} +if (lean_is_scalar(x_406)) { + x_407 = lean_alloc_ctor(0, 2, 0); +} else { + x_407 = x_406; +} +lean_ctor_set(x_407, 0, x_1); +lean_ctor_set(x_407, 1, x_405); +x_408 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_408, 0, x_407); +lean_ctor_set(x_408, 1, x_404); +return x_408; } } else { -uint8_t x_269; -lean_dec(x_243); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +uint8_t x_409; lean_dec(x_1); -x_269 = !lean_is_exclusive(x_244); -if (x_269 == 0) -{ -return x_244; -} -else +x_409 = !lean_is_exclusive(x_393); +if (x_409 == 0) { -lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_270 = lean_ctor_get(x_244, 0); -x_271 = lean_ctor_get(x_244, 1); -lean_inc(x_271); -lean_inc(x_270); -lean_dec(x_244); -x_272 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_272, 0, x_270); -lean_ctor_set(x_272, 1, x_271); -return x_272; +lean_object* x_410; uint8_t x_411; +x_410 = lean_ctor_get(x_393, 0); +lean_dec(x_410); +x_411 = !lean_is_exclusive(x_394); +if (x_411 == 0) +{ +lean_object* x_412; lean_object* x_413; lean_object* x_414; +x_412 = lean_ctor_get(x_394, 0); +lean_dec(x_412); +x_413 = lean_ctor_get(x_395, 0); +lean_inc(x_413); +lean_dec(x_395); +x_414 = l_Lean_mkAppN(x_2, x_413); +lean_dec(x_413); +lean_ctor_set(x_394, 0, x_414); +return x_393; +} +else +{ +lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; +x_415 = lean_ctor_get(x_394, 1); +lean_inc(x_415); +lean_dec(x_394); +x_416 = lean_ctor_get(x_395, 0); +lean_inc(x_416); +lean_dec(x_395); +x_417 = l_Lean_mkAppN(x_2, x_416); +lean_dec(x_416); +x_418 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_418, 0, x_417); +lean_ctor_set(x_418, 1, x_415); +lean_ctor_set(x_393, 0, x_418); +return x_393; +} +} +else +{ +lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; +x_419 = lean_ctor_get(x_393, 1); +lean_inc(x_419); +lean_dec(x_393); +x_420 = lean_ctor_get(x_394, 1); +lean_inc(x_420); +if (lean_is_exclusive(x_394)) { + lean_ctor_release(x_394, 0); + lean_ctor_release(x_394, 1); + x_421 = x_394; +} else { + lean_dec_ref(x_394); + x_421 = lean_box(0); } +x_422 = lean_ctor_get(x_395, 0); +lean_inc(x_422); +lean_dec(x_395); +x_423 = l_Lean_mkAppN(x_2, x_422); +lean_dec(x_422); +if (lean_is_scalar(x_421)) { + x_424 = lean_alloc_ctor(0, 2, 0); +} else { + x_424 = x_421; } +lean_ctor_set(x_424, 0, x_423); +lean_ctor_set(x_424, 1, x_420); +x_425 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_425, 0, x_424); +lean_ctor_set(x_425, 1, x_419); +return x_425; } -block_307: -{ -lean_object* x_275; -lean_dec(x_274); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_275 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_275) == 0) -{ -lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; -x_276 = lean_ctor_get(x_275, 0); -lean_inc(x_276); -x_277 = lean_ctor_get(x_275, 1); -lean_inc(x_277); -lean_dec(x_275); -x_278 = lean_ctor_get(x_276, 0); -lean_inc(x_278); -lean_dec(x_276); -x_279 = lean_array_get_size(x_3); -x_280 = lean_unsigned_to_nat(0u); -x_281 = lean_unsigned_to_nat(1u); -x_282 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_282, 0, x_280); -lean_ctor_set(x_282, 1, x_279); -lean_ctor_set(x_282, 2, x_281); -x_283 = 0; -x_284 = lean_box(x_283); -x_285 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_285, 0, x_3); -lean_ctor_set(x_285, 1, x_284); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_286 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_278, x_282, x_282, x_285, x_280, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_277); -lean_dec(x_282); -lean_dec(x_278); -if (lean_obj_tag(x_286) == 0) -{ -lean_object* x_287; lean_object* x_288; lean_object* x_289; uint8_t x_290; -x_287 = lean_ctor_get(x_286, 0); -lean_inc(x_287); -x_288 = lean_ctor_get(x_287, 0); -lean_inc(x_288); -x_289 = lean_ctor_get(x_288, 1); -lean_inc(x_289); -x_290 = lean_unbox(x_289); -lean_dec(x_289); -if (x_290 == 0) -{ -lean_object* x_291; lean_object* x_292; lean_object* x_293; -lean_dec(x_288); -lean_dec(x_2); -x_291 = lean_ctor_get(x_286, 1); -lean_inc(x_291); -lean_dec(x_286); -x_292 = lean_ctor_get(x_287, 1); -lean_inc(x_292); -lean_dec(x_287); -lean_inc(x_1); -x_293 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_292, x_7, x_8, x_9, x_10, x_291); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_293; -} -else -{ -lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_294 = lean_ctor_get(x_286, 1); -lean_inc(x_294); -lean_dec(x_286); -x_295 = lean_ctor_get(x_287, 1); -lean_inc(x_295); -lean_dec(x_287); -x_296 = lean_ctor_get(x_288, 0); -lean_inc(x_296); -lean_dec(x_288); -x_297 = l_Lean_mkAppN(x_2, x_296); -lean_dec(x_296); -x_298 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_297, x_5, x_295, x_7, x_8, x_9, x_10, x_294); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_298; } } else { -uint8_t x_299; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +uint8_t x_426; lean_dec(x_2); lean_dec(x_1); -x_299 = !lean_is_exclusive(x_286); -if (x_299 == 0) +x_426 = !lean_is_exclusive(x_393); +if (x_426 == 0) { -return x_286; +return x_393; } else { -lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_300 = lean_ctor_get(x_286, 0); -x_301 = lean_ctor_get(x_286, 1); -lean_inc(x_301); -lean_inc(x_300); -lean_dec(x_286); -x_302 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_302, 0, x_300); -lean_ctor_set(x_302, 1, x_301); -return x_302; +lean_object* x_427; lean_object* x_428; lean_object* x_429; +x_427 = lean_ctor_get(x_393, 0); +x_428 = lean_ctor_get(x_393, 1); +lean_inc(x_428); +lean_inc(x_427); +lean_dec(x_393); +x_429 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_429, 0, x_427); +lean_ctor_set(x_429, 1, x_428); +return x_429; } } } else { -uint8_t x_303; +uint8_t x_430; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8210,710 +8207,620 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_303 = !lean_is_exclusive(x_275); -if (x_303 == 0) +x_430 = !lean_is_exclusive(x_382); +if (x_430 == 0) { -return x_275; +return x_382; } else { -lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_304 = lean_ctor_get(x_275, 0); -x_305 = lean_ctor_get(x_275, 1); -lean_inc(x_305); -lean_inc(x_304); -lean_dec(x_275); -x_306 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_306, 0, x_304); -lean_ctor_set(x_306, 1, x_305); -return x_306; +lean_object* x_431; lean_object* x_432; lean_object* x_433; +x_431 = lean_ctor_get(x_382, 0); +x_432 = lean_ctor_get(x_382, 1); +lean_inc(x_432); +lean_inc(x_431); +lean_dec(x_382); +x_433 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_433, 0, x_431); +lean_ctor_set(x_433, 1, x_432); +return x_433; } } } } -case 4: +case 3: { -lean_object* x_320; lean_object* x_351; lean_object* x_385; uint8_t x_386; +lean_object* x_447; lean_object* x_526; lean_object* x_580; uint8_t x_581; lean_dec(x_4); -x_385 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_386 = l_Lean_Expr_isConstOf(x_2, x_385); -if (x_386 == 0) +x_580 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_581 = l_Lean_Expr_isConstOf(x_2, x_580); +if (x_581 == 0) { -lean_object* x_387; -x_387 = lean_box(0); -x_351 = x_387; -goto block_384; +lean_object* x_582; +x_582 = lean_box(0); +x_526 = x_582; +goto block_579; } else { -lean_object* x_388; lean_object* x_389; uint8_t x_390; -x_388 = lean_array_get_size(x_3); -x_389 = lean_unsigned_to_nat(2u); -x_390 = lean_nat_dec_eq(x_388, x_389); -if (x_390 == 0) +lean_object* x_583; lean_object* x_584; uint8_t x_585; +x_583 = lean_array_get_size(x_3); +x_584 = lean_unsigned_to_nat(2u); +x_585 = lean_nat_dec_eq(x_583, x_584); +if (x_585 == 0) { -lean_object* x_391; -lean_dec(x_388); -x_391 = lean_box(0); -x_351 = x_391; -goto block_384; +lean_object* x_586; +lean_dec(x_583); +x_586 = lean_box(0); +x_526 = x_586; +goto block_579; } else { -lean_object* x_392; uint8_t x_393; -x_392 = lean_unsigned_to_nat(0u); -x_393 = lean_nat_dec_lt(x_392, x_388); -lean_dec(x_388); -if (x_393 == 0) +lean_object* x_587; uint8_t x_588; +x_587 = lean_unsigned_to_nat(0u); +x_588 = lean_nat_dec_lt(x_587, x_583); +lean_dec(x_583); +if (x_588 == 0) { -lean_object* x_394; lean_object* x_395; -x_394 = l_Lean_instInhabitedExpr; -x_395 = l_outOfBounds___rarg(x_394); -x_320 = x_395; -goto block_350; +lean_object* x_589; lean_object* x_590; +x_589 = l_Lean_instInhabitedExpr; +x_590 = l_outOfBounds___rarg(x_589); +x_447 = x_590; +goto block_525; } else { -lean_object* x_396; -x_396 = lean_array_fget(x_3, x_392); -x_320 = x_396; -goto block_350; +lean_object* x_591; +x_591 = lean_array_fget(x_3, x_587); +x_447 = x_591; +goto block_525; } } } -block_350: +block_525: { -lean_object* x_321; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_320); -x_321 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_320, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_321) == 0) +lean_object* x_448; +lean_inc(x_447); +x_448 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_447, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_448) == 0) +{ +uint8_t x_449; +x_449 = !lean_is_exclusive(x_448); +if (x_449 == 0) +{ +lean_object* x_450; uint8_t x_451; +x_450 = lean_ctor_get(x_448, 0); +x_451 = !lean_is_exclusive(x_450); +if (x_451 == 0) +{ +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; +x_452 = lean_ctor_get(x_450, 0); +x_453 = lean_ctor_get(x_450, 1); +x_454 = lean_ctor_get(x_453, 2); +lean_inc(x_454); +lean_inc(x_454); +x_455 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_454, x_452); +if (lean_obj_tag(x_455) == 0) { -lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; -x_322 = lean_ctor_get(x_321, 0); -lean_inc(x_322); -x_323 = lean_ctor_get(x_321, 1); -lean_inc(x_323); -lean_dec(x_321); -x_324 = lean_ctor_get(x_322, 0); -lean_inc(x_324); -x_325 = lean_ctor_get(x_322, 1); -lean_inc(x_325); -lean_dec(x_322); -x_326 = lean_ctor_get(x_325, 2); -lean_inc(x_326); -lean_inc(x_326); -x_327 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_326, x_324); -if (lean_obj_tag(x_327) == 0) -{ -size_t x_328; size_t x_329; uint8_t x_330; -x_328 = lean_ptr_addr(x_320); -lean_dec(x_320); -x_329 = lean_ptr_addr(x_324); -x_330 = lean_usize_dec_eq(x_328, x_329); -if (x_330 == 0) -{ -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; -x_331 = lean_unsigned_to_nat(0u); -lean_inc(x_324); -x_332 = lean_array_set(x_3, x_331, x_324); -x_333 = l_Lean_mkAppN(x_2, x_332); -lean_dec(x_332); -x_334 = lean_ctor_get(x_325, 0); -lean_inc(x_334); -x_335 = lean_ctor_get(x_325, 1); -lean_inc(x_335); -lean_dec(x_325); -lean_inc(x_333); -x_336 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_326, x_324, x_333); -x_337 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_337, 0, x_334); -lean_ctor_set(x_337, 1, x_335); -lean_ctor_set(x_337, 2, x_336); -x_338 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_333, x_5, x_337, x_7, x_8, x_9, x_10, x_323); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_338; +size_t x_456; size_t x_457; uint8_t x_458; +x_456 = lean_ptr_addr(x_447); +lean_dec(x_447); +x_457 = lean_ptr_addr(x_452); +x_458 = lean_usize_dec_eq(x_456, x_457); +if (x_458 == 0) +{ +lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; +lean_dec(x_1); +x_459 = lean_unsigned_to_nat(0u); +lean_inc(x_452); +x_460 = lean_array_set(x_3, x_459, x_452); +x_461 = l_Lean_mkAppN(x_2, x_460); +lean_dec(x_460); +x_462 = lean_ctor_get(x_453, 0); +lean_inc(x_462); +x_463 = lean_ctor_get(x_453, 1); +lean_inc(x_463); +lean_dec(x_453); +lean_inc(x_461); +x_464 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_454, x_452, x_461); +x_465 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_465, 0, x_462); +lean_ctor_set(x_465, 1, x_463); +lean_ctor_set(x_465, 2, x_464); +lean_ctor_set(x_450, 1, x_465); +lean_ctor_set(x_450, 0, x_461); +return x_448; } else { -lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; +lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_dec(x_3); lean_dec(x_2); -x_339 = lean_ctor_get(x_325, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_325, 1); -lean_inc(x_340); -lean_dec(x_325); -lean_inc(x_1); -x_341 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_326, x_324, x_1); -x_342 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_342, 0, x_339); -lean_ctor_set(x_342, 1, x_340); -lean_ctor_set(x_342, 2, x_341); +x_466 = lean_ctor_get(x_453, 0); +lean_inc(x_466); +x_467 = lean_ctor_get(x_453, 1); +lean_inc(x_467); +lean_dec(x_453); lean_inc(x_1); -x_343 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_342, x_7, x_8, x_9, x_10, x_323); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_343; +x_468 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_454, x_452, x_1); +x_469 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_469, 0, x_466); +lean_ctor_set(x_469, 1, x_467); +lean_ctor_set(x_469, 2, x_468); +lean_ctor_set(x_450, 1, x_469); +lean_ctor_set(x_450, 0, x_1); +return x_448; } } else { -lean_object* x_344; lean_object* x_345; -lean_dec(x_326); -lean_dec(x_324); -lean_dec(x_320); +lean_object* x_470; +lean_dec(x_454); +lean_dec(x_452); +lean_dec(x_447); lean_dec(x_3); lean_dec(x_2); -x_344 = lean_ctor_get(x_327, 0); -lean_inc(x_344); -lean_dec(x_327); -x_345 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_344, x_5, x_325, x_7, x_8, x_9, x_10, x_323); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_345; +lean_dec(x_1); +x_470 = lean_ctor_get(x_455, 0); +lean_inc(x_470); +lean_dec(x_455); +lean_ctor_set(x_450, 0, x_470); +return x_448; } } else { -uint8_t x_346; -lean_dec(x_320); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; +x_471 = lean_ctor_get(x_450, 0); +x_472 = lean_ctor_get(x_450, 1); +lean_inc(x_472); +lean_inc(x_471); +lean_dec(x_450); +x_473 = lean_ctor_get(x_472, 2); +lean_inc(x_473); +lean_inc(x_473); +x_474 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_473, x_471); +if (lean_obj_tag(x_474) == 0) +{ +size_t x_475; size_t x_476; uint8_t x_477; +x_475 = lean_ptr_addr(x_447); +lean_dec(x_447); +x_476 = lean_ptr_addr(x_471); +x_477 = lean_usize_dec_eq(x_475, x_476); +if (x_477 == 0) +{ +lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_dec(x_1); -x_346 = !lean_is_exclusive(x_321); -if (x_346 == 0) +x_478 = lean_unsigned_to_nat(0u); +lean_inc(x_471); +x_479 = lean_array_set(x_3, x_478, x_471); +x_480 = l_Lean_mkAppN(x_2, x_479); +lean_dec(x_479); +x_481 = lean_ctor_get(x_472, 0); +lean_inc(x_481); +x_482 = lean_ctor_get(x_472, 1); +lean_inc(x_482); +lean_dec(x_472); +lean_inc(x_480); +x_483 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_473, x_471, x_480); +x_484 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_484, 0, x_481); +lean_ctor_set(x_484, 1, x_482); +lean_ctor_set(x_484, 2, x_483); +x_485 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_485, 0, x_480); +lean_ctor_set(x_485, 1, x_484); +lean_ctor_set(x_448, 0, x_485); +return x_448; +} +else { -return x_321; +lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; +lean_dec(x_3); +lean_dec(x_2); +x_486 = lean_ctor_get(x_472, 0); +lean_inc(x_486); +x_487 = lean_ctor_get(x_472, 1); +lean_inc(x_487); +lean_dec(x_472); +lean_inc(x_1); +x_488 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_473, x_471, x_1); +x_489 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_489, 0, x_486); +lean_ctor_set(x_489, 1, x_487); +lean_ctor_set(x_489, 2, x_488); +x_490 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_490, 0, x_1); +lean_ctor_set(x_490, 1, x_489); +lean_ctor_set(x_448, 0, x_490); +return x_448; +} } else { -lean_object* x_347; lean_object* x_348; lean_object* x_349; -x_347 = lean_ctor_get(x_321, 0); -x_348 = lean_ctor_get(x_321, 1); -lean_inc(x_348); -lean_inc(x_347); -lean_dec(x_321); -x_349 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_349, 0, x_347); -lean_ctor_set(x_349, 1, x_348); -return x_349; +lean_object* x_491; lean_object* x_492; +lean_dec(x_473); +lean_dec(x_471); +lean_dec(x_447); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_491 = lean_ctor_get(x_474, 0); +lean_inc(x_491); +lean_dec(x_474); +x_492 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_492, 0, x_491); +lean_ctor_set(x_492, 1, x_472); +lean_ctor_set(x_448, 0, x_492); +return x_448; } } } -block_384: -{ -lean_object* x_352; -lean_dec(x_351); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_352 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_352) == 0) -{ -lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; uint8_t x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; -x_353 = lean_ctor_get(x_352, 0); -lean_inc(x_353); -x_354 = lean_ctor_get(x_352, 1); -lean_inc(x_354); -lean_dec(x_352); -x_355 = lean_ctor_get(x_353, 0); -lean_inc(x_355); -lean_dec(x_353); -x_356 = lean_array_get_size(x_3); -x_357 = lean_unsigned_to_nat(0u); -x_358 = lean_unsigned_to_nat(1u); -x_359 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_359, 0, x_357); -lean_ctor_set(x_359, 1, x_356); -lean_ctor_set(x_359, 2, x_358); -x_360 = 0; -x_361 = lean_box(x_360); -x_362 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_362, 0, x_3); -lean_ctor_set(x_362, 1, x_361); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_363 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_355, x_359, x_359, x_362, x_357, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_354); -lean_dec(x_359); -lean_dec(x_355); -if (lean_obj_tag(x_363) == 0) -{ -lean_object* x_364; lean_object* x_365; lean_object* x_366; uint8_t x_367; -x_364 = lean_ctor_get(x_363, 0); -lean_inc(x_364); -x_365 = lean_ctor_get(x_364, 0); -lean_inc(x_365); -x_366 = lean_ctor_get(x_365, 1); -lean_inc(x_366); -x_367 = lean_unbox(x_366); -lean_dec(x_366); -if (x_367 == 0) -{ -lean_object* x_368; lean_object* x_369; lean_object* x_370; -lean_dec(x_365); -lean_dec(x_2); -x_368 = lean_ctor_get(x_363, 1); -lean_inc(x_368); -lean_dec(x_363); -x_369 = lean_ctor_get(x_364, 1); -lean_inc(x_369); -lean_dec(x_364); -lean_inc(x_1); -x_370 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_369, x_7, x_8, x_9, x_10, x_368); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_370; -} else { -lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -x_371 = lean_ctor_get(x_363, 1); -lean_inc(x_371); -lean_dec(x_363); -x_372 = lean_ctor_get(x_364, 1); -lean_inc(x_372); -lean_dec(x_364); -x_373 = lean_ctor_get(x_365, 0); -lean_inc(x_373); -lean_dec(x_365); -x_374 = l_Lean_mkAppN(x_2, x_373); -lean_dec(x_373); -x_375 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_374, x_5, x_372, x_7, x_8, x_9, x_10, x_371); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_375; -} +lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; +x_493 = lean_ctor_get(x_448, 0); +x_494 = lean_ctor_get(x_448, 1); +lean_inc(x_494); +lean_inc(x_493); +lean_dec(x_448); +x_495 = lean_ctor_get(x_493, 0); +lean_inc(x_495); +x_496 = lean_ctor_get(x_493, 1); +lean_inc(x_496); +if (lean_is_exclusive(x_493)) { + lean_ctor_release(x_493, 0); + lean_ctor_release(x_493, 1); + x_497 = x_493; +} else { + lean_dec_ref(x_493); + x_497 = lean_box(0); } -else +x_498 = lean_ctor_get(x_496, 2); +lean_inc(x_498); +lean_inc(x_498); +x_499 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_498, x_495); +if (lean_obj_tag(x_499) == 0) { -uint8_t x_376; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_376 = !lean_is_exclusive(x_363); -if (x_376 == 0) +size_t x_500; size_t x_501; uint8_t x_502; +x_500 = lean_ptr_addr(x_447); +lean_dec(x_447); +x_501 = lean_ptr_addr(x_495); +x_502 = lean_usize_dec_eq(x_500, x_501); +if (x_502 == 0) { -return x_363; +lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; +lean_dec(x_1); +x_503 = lean_unsigned_to_nat(0u); +lean_inc(x_495); +x_504 = lean_array_set(x_3, x_503, x_495); +x_505 = l_Lean_mkAppN(x_2, x_504); +lean_dec(x_504); +x_506 = lean_ctor_get(x_496, 0); +lean_inc(x_506); +x_507 = lean_ctor_get(x_496, 1); +lean_inc(x_507); +lean_dec(x_496); +lean_inc(x_505); +x_508 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_498, x_495, x_505); +x_509 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_509, 0, x_506); +lean_ctor_set(x_509, 1, x_507); +lean_ctor_set(x_509, 2, x_508); +if (lean_is_scalar(x_497)) { + x_510 = lean_alloc_ctor(0, 2, 0); +} else { + x_510 = x_497; +} +lean_ctor_set(x_510, 0, x_505); +lean_ctor_set(x_510, 1, x_509); +x_511 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_511, 0, x_510); +lean_ctor_set(x_511, 1, x_494); +return x_511; } else { -lean_object* x_377; lean_object* x_378; lean_object* x_379; -x_377 = lean_ctor_get(x_363, 0); -x_378 = lean_ctor_get(x_363, 1); -lean_inc(x_378); -lean_inc(x_377); -lean_dec(x_363); -x_379 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_379, 0, x_377); -lean_ctor_set(x_379, 1, x_378); -return x_379; +lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; +lean_dec(x_3); +lean_dec(x_2); +x_512 = lean_ctor_get(x_496, 0); +lean_inc(x_512); +x_513 = lean_ctor_get(x_496, 1); +lean_inc(x_513); +lean_dec(x_496); +lean_inc(x_1); +x_514 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_498, x_495, x_1); +x_515 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_515, 0, x_512); +lean_ctor_set(x_515, 1, x_513); +lean_ctor_set(x_515, 2, x_514); +if (lean_is_scalar(x_497)) { + x_516 = lean_alloc_ctor(0, 2, 0); +} else { + x_516 = x_497; } +lean_ctor_set(x_516, 0, x_1); +lean_ctor_set(x_516, 1, x_515); +x_517 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_517, 0, x_516); +lean_ctor_set(x_517, 1, x_494); +return x_517; } } else { -uint8_t x_380; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_518; lean_object* x_519; lean_object* x_520; +lean_dec(x_498); +lean_dec(x_495); +lean_dec(x_447); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_380 = !lean_is_exclusive(x_352); -if (x_380 == 0) -{ -return x_352; -} -else -{ -lean_object* x_381; lean_object* x_382; lean_object* x_383; -x_381 = lean_ctor_get(x_352, 0); -x_382 = lean_ctor_get(x_352, 1); -lean_inc(x_382); -lean_inc(x_381); -lean_dec(x_352); -x_383 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_383, 0, x_381); -lean_ctor_set(x_383, 1, x_382); -return x_383; -} -} -} +x_518 = lean_ctor_get(x_499, 0); +lean_inc(x_518); +lean_dec(x_499); +if (lean_is_scalar(x_497)) { + x_519 = lean_alloc_ctor(0, 2, 0); +} else { + x_519 = x_497; } -case 5: -{ -lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; -x_397 = lean_ctor_get(x_2, 0); -lean_inc(x_397); -x_398 = lean_ctor_get(x_2, 1); -lean_inc(x_398); -lean_dec(x_2); -x_399 = lean_array_set(x_3, x_4, x_398); -x_400 = lean_unsigned_to_nat(1u); -x_401 = lean_nat_sub(x_4, x_400); -lean_dec(x_4); -x_2 = x_397; -x_3 = x_399; -x_4 = x_401; -goto _start; +lean_ctor_set(x_519, 0, x_518); +lean_ctor_set(x_519, 1, x_496); +x_520 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_520, 0, x_519); +lean_ctor_set(x_520, 1, x_494); +return x_520; } -case 6: -{ -lean_object* x_403; lean_object* x_434; lean_object* x_468; uint8_t x_469; -lean_dec(x_4); -x_468 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_469 = l_Lean_Expr_isConstOf(x_2, x_468); -if (x_469 == 0) -{ -lean_object* x_470; -x_470 = lean_box(0); -x_434 = x_470; -goto block_467; } -else -{ -lean_object* x_471; lean_object* x_472; uint8_t x_473; -x_471 = lean_array_get_size(x_3); -x_472 = lean_unsigned_to_nat(2u); -x_473 = lean_nat_dec_eq(x_471, x_472); -if (x_473 == 0) -{ -lean_object* x_474; -lean_dec(x_471); -x_474 = lean_box(0); -x_434 = x_474; -goto block_467; } else { -lean_object* x_475; uint8_t x_476; -x_475 = lean_unsigned_to_nat(0u); -x_476 = lean_nat_dec_lt(x_475, x_471); -lean_dec(x_471); -if (x_476 == 0) +uint8_t x_521; +lean_dec(x_447); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_521 = !lean_is_exclusive(x_448); +if (x_521 == 0) { -lean_object* x_477; lean_object* x_478; -x_477 = l_Lean_instInhabitedExpr; -x_478 = l_outOfBounds___rarg(x_477); -x_403 = x_478; -goto block_433; +return x_448; } else { -lean_object* x_479; -x_479 = lean_array_fget(x_3, x_475); -x_403 = x_479; -goto block_433; +lean_object* x_522; lean_object* x_523; lean_object* x_524; +x_522 = lean_ctor_get(x_448, 0); +x_523 = lean_ctor_get(x_448, 1); +lean_inc(x_523); +lean_inc(x_522); +lean_dec(x_448); +x_524 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_524, 0, x_522); +lean_ctor_set(x_524, 1, x_523); +return x_524; } } } -block_433: +block_579: { -lean_object* x_404; +lean_object* x_527; +lean_dec(x_526); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_403); -x_404 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_403, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_404) == 0) -{ -lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; -x_405 = lean_ctor_get(x_404, 0); -lean_inc(x_405); -x_406 = lean_ctor_get(x_404, 1); -lean_inc(x_406); -lean_dec(x_404); -x_407 = lean_ctor_get(x_405, 0); -lean_inc(x_407); -x_408 = lean_ctor_get(x_405, 1); -lean_inc(x_408); -lean_dec(x_405); -x_409 = lean_ctor_get(x_408, 2); -lean_inc(x_409); -lean_inc(x_409); -x_410 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_409, x_407); -if (lean_obj_tag(x_410) == 0) -{ -size_t x_411; size_t x_412; uint8_t x_413; -x_411 = lean_ptr_addr(x_403); -lean_dec(x_403); -x_412 = lean_ptr_addr(x_407); -x_413 = lean_usize_dec_eq(x_411, x_412); -if (x_413 == 0) -{ -lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; -x_414 = lean_unsigned_to_nat(0u); -lean_inc(x_407); -x_415 = lean_array_set(x_3, x_414, x_407); -x_416 = l_Lean_mkAppN(x_2, x_415); -lean_dec(x_415); -x_417 = lean_ctor_get(x_408, 0); -lean_inc(x_417); -x_418 = lean_ctor_get(x_408, 1); -lean_inc(x_418); -lean_dec(x_408); -lean_inc(x_416); -x_419 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_409, x_407, x_416); -x_420 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_420, 0, x_417); -lean_ctor_set(x_420, 1, x_418); -lean_ctor_set(x_420, 2, x_419); -x_421 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_416, x_5, x_420, x_7, x_8, x_9, x_10, x_406); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_421; -} -else +lean_inc(x_2); +x_527 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_527) == 0) { -lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; +lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; uint8_t x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; +x_528 = lean_ctor_get(x_527, 0); +lean_inc(x_528); +x_529 = lean_ctor_get(x_527, 1); +lean_inc(x_529); +lean_dec(x_527); +x_530 = lean_ctor_get(x_528, 0); +lean_inc(x_530); +lean_dec(x_528); +x_531 = lean_array_get_size(x_3); +x_532 = lean_unsigned_to_nat(0u); +x_533 = lean_unsigned_to_nat(1u); +lean_inc(x_531); +x_534 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_534, 0, x_532); +lean_ctor_set(x_534, 1, x_531); +lean_ctor_set(x_534, 2, x_533); +x_535 = 0; +x_536 = lean_box(x_535); +lean_inc(x_3); +x_537 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_537, 0, x_3); +lean_ctor_set(x_537, 1, x_536); +lean_inc(x_2); +x_538 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_530, x_531, x_534, x_534, x_537, x_532, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_529); +lean_dec(x_534); +lean_dec(x_531); +lean_dec(x_530); lean_dec(x_3); +if (lean_obj_tag(x_538) == 0) +{ +lean_object* x_539; lean_object* x_540; lean_object* x_541; uint8_t x_542; +x_539 = lean_ctor_get(x_538, 0); +lean_inc(x_539); +x_540 = lean_ctor_get(x_539, 0); +lean_inc(x_540); +x_541 = lean_ctor_get(x_540, 1); +lean_inc(x_541); +x_542 = lean_unbox(x_541); +lean_dec(x_541); +if (x_542 == 0) +{ +uint8_t x_543; +lean_dec(x_540); lean_dec(x_2); -x_422 = lean_ctor_get(x_408, 0); -lean_inc(x_422); -x_423 = lean_ctor_get(x_408, 1); -lean_inc(x_423); -lean_dec(x_408); -lean_inc(x_1); -x_424 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_409, x_407, x_1); -x_425 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_425, 0, x_422); -lean_ctor_set(x_425, 1, x_423); -lean_ctor_set(x_425, 2, x_424); -lean_inc(x_1); -x_426 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_425, x_7, x_8, x_9, x_10, x_406); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_426; +x_543 = !lean_is_exclusive(x_538); +if (x_543 == 0) +{ +lean_object* x_544; uint8_t x_545; +x_544 = lean_ctor_get(x_538, 0); +lean_dec(x_544); +x_545 = !lean_is_exclusive(x_539); +if (x_545 == 0) +{ +lean_object* x_546; +x_546 = lean_ctor_get(x_539, 0); +lean_dec(x_546); +lean_ctor_set(x_539, 0, x_1); +return x_538; +} +else +{ +lean_object* x_547; lean_object* x_548; +x_547 = lean_ctor_get(x_539, 1); +lean_inc(x_547); +lean_dec(x_539); +x_548 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_548, 0, x_1); +lean_ctor_set(x_548, 1, x_547); +lean_ctor_set(x_538, 0, x_548); +return x_538; +} +} +else +{ +lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; +x_549 = lean_ctor_get(x_538, 1); +lean_inc(x_549); +lean_dec(x_538); +x_550 = lean_ctor_get(x_539, 1); +lean_inc(x_550); +if (lean_is_exclusive(x_539)) { + lean_ctor_release(x_539, 0); + lean_ctor_release(x_539, 1); + x_551 = x_539; +} else { + lean_dec_ref(x_539); + x_551 = lean_box(0); } +if (lean_is_scalar(x_551)) { + x_552 = lean_alloc_ctor(0, 2, 0); +} else { + x_552 = x_551; } -else -{ -lean_object* x_427; lean_object* x_428; -lean_dec(x_409); -lean_dec(x_407); -lean_dec(x_403); -lean_dec(x_3); -lean_dec(x_2); -x_427 = lean_ctor_get(x_410, 0); -lean_inc(x_427); -lean_dec(x_410); -x_428 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_427, x_5, x_408, x_7, x_8, x_9, x_10, x_406); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_428; +lean_ctor_set(x_552, 0, x_1); +lean_ctor_set(x_552, 1, x_550); +x_553 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_553, 0, x_552); +lean_ctor_set(x_553, 1, x_549); +return x_553; } } else { -uint8_t x_429; -lean_dec(x_403); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +uint8_t x_554; lean_dec(x_1); -x_429 = !lean_is_exclusive(x_404); -if (x_429 == 0) +x_554 = !lean_is_exclusive(x_538); +if (x_554 == 0) +{ +lean_object* x_555; uint8_t x_556; +x_555 = lean_ctor_get(x_538, 0); +lean_dec(x_555); +x_556 = !lean_is_exclusive(x_539); +if (x_556 == 0) { -return x_404; +lean_object* x_557; lean_object* x_558; lean_object* x_559; +x_557 = lean_ctor_get(x_539, 0); +lean_dec(x_557); +x_558 = lean_ctor_get(x_540, 0); +lean_inc(x_558); +lean_dec(x_540); +x_559 = l_Lean_mkAppN(x_2, x_558); +lean_dec(x_558); +lean_ctor_set(x_539, 0, x_559); +return x_538; } else { -lean_object* x_430; lean_object* x_431; lean_object* x_432; -x_430 = lean_ctor_get(x_404, 0); -x_431 = lean_ctor_get(x_404, 1); -lean_inc(x_431); -lean_inc(x_430); -lean_dec(x_404); -x_432 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_432, 0, x_430); -lean_ctor_set(x_432, 1, x_431); -return x_432; -} -} +lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; +x_560 = lean_ctor_get(x_539, 1); +lean_inc(x_560); +lean_dec(x_539); +x_561 = lean_ctor_get(x_540, 0); +lean_inc(x_561); +lean_dec(x_540); +x_562 = l_Lean_mkAppN(x_2, x_561); +lean_dec(x_561); +x_563 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_563, 0, x_562); +lean_ctor_set(x_563, 1, x_560); +lean_ctor_set(x_538, 0, x_563); +return x_538; } -block_467: -{ -lean_object* x_435; -lean_dec(x_434); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_435 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_435) == 0) -{ -lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; uint8_t x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; -x_436 = lean_ctor_get(x_435, 0); -lean_inc(x_436); -x_437 = lean_ctor_get(x_435, 1); -lean_inc(x_437); -lean_dec(x_435); -x_438 = lean_ctor_get(x_436, 0); -lean_inc(x_438); -lean_dec(x_436); -x_439 = lean_array_get_size(x_3); -x_440 = lean_unsigned_to_nat(0u); -x_441 = lean_unsigned_to_nat(1u); -x_442 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_442, 0, x_440); -lean_ctor_set(x_442, 1, x_439); -lean_ctor_set(x_442, 2, x_441); -x_443 = 0; -x_444 = lean_box(x_443); -x_445 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_445, 0, x_3); -lean_ctor_set(x_445, 1, x_444); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_446 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_438, x_442, x_442, x_445, x_440, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_437); -lean_dec(x_442); -lean_dec(x_438); -if (lean_obj_tag(x_446) == 0) -{ -lean_object* x_447; lean_object* x_448; lean_object* x_449; uint8_t x_450; -x_447 = lean_ctor_get(x_446, 0); -lean_inc(x_447); -x_448 = lean_ctor_get(x_447, 0); -lean_inc(x_448); -x_449 = lean_ctor_get(x_448, 1); -lean_inc(x_449); -x_450 = lean_unbox(x_449); -lean_dec(x_449); -if (x_450 == 0) -{ -lean_object* x_451; lean_object* x_452; lean_object* x_453; -lean_dec(x_448); -lean_dec(x_2); -x_451 = lean_ctor_get(x_446, 1); -lean_inc(x_451); -lean_dec(x_446); -x_452 = lean_ctor_get(x_447, 1); -lean_inc(x_452); -lean_dec(x_447); -lean_inc(x_1); -x_453 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_452, x_7, x_8, x_9, x_10, x_451); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_453; } else { -lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; -x_454 = lean_ctor_get(x_446, 1); -lean_inc(x_454); -lean_dec(x_446); -x_455 = lean_ctor_get(x_447, 1); -lean_inc(x_455); -lean_dec(x_447); -x_456 = lean_ctor_get(x_448, 0); -lean_inc(x_456); -lean_dec(x_448); -x_457 = l_Lean_mkAppN(x_2, x_456); -lean_dec(x_456); -x_458 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_457, x_5, x_455, x_7, x_8, x_9, x_10, x_454); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_458; +lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; +x_564 = lean_ctor_get(x_538, 1); +lean_inc(x_564); +lean_dec(x_538); +x_565 = lean_ctor_get(x_539, 1); +lean_inc(x_565); +if (lean_is_exclusive(x_539)) { + lean_ctor_release(x_539, 0); + lean_ctor_release(x_539, 1); + x_566 = x_539; +} else { + lean_dec_ref(x_539); + x_566 = lean_box(0); +} +x_567 = lean_ctor_get(x_540, 0); +lean_inc(x_567); +lean_dec(x_540); +x_568 = l_Lean_mkAppN(x_2, x_567); +lean_dec(x_567); +if (lean_is_scalar(x_566)) { + x_569 = lean_alloc_ctor(0, 2, 0); +} else { + x_569 = x_566; +} +lean_ctor_set(x_569, 0, x_568); +lean_ctor_set(x_569, 1, x_565); +x_570 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_570, 0, x_569); +lean_ctor_set(x_570, 1, x_564); +return x_570; +} } } else { -uint8_t x_459; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +uint8_t x_571; lean_dec(x_2); lean_dec(x_1); -x_459 = !lean_is_exclusive(x_446); -if (x_459 == 0) +x_571 = !lean_is_exclusive(x_538); +if (x_571 == 0) { -return x_446; +return x_538; } else { -lean_object* x_460; lean_object* x_461; lean_object* x_462; -x_460 = lean_ctor_get(x_446, 0); -x_461 = lean_ctor_get(x_446, 1); -lean_inc(x_461); -lean_inc(x_460); -lean_dec(x_446); -x_462 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_462, 0, x_460); -lean_ctor_set(x_462, 1, x_461); -return x_462; +lean_object* x_572; lean_object* x_573; lean_object* x_574; +x_572 = lean_ctor_get(x_538, 0); +x_573 = lean_ctor_get(x_538, 1); +lean_inc(x_573); +lean_inc(x_572); +lean_dec(x_538); +x_574 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_574, 0, x_572); +lean_ctor_set(x_574, 1, x_573); +return x_574; } } } else { -uint8_t x_463; +uint8_t x_575; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8923,693 +8830,620 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_463 = !lean_is_exclusive(x_435); -if (x_463 == 0) +x_575 = !lean_is_exclusive(x_527); +if (x_575 == 0) { -return x_435; +return x_527; } else { -lean_object* x_464; lean_object* x_465; lean_object* x_466; -x_464 = lean_ctor_get(x_435, 0); -x_465 = lean_ctor_get(x_435, 1); -lean_inc(x_465); -lean_inc(x_464); -lean_dec(x_435); -x_466 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_466, 0, x_464); -lean_ctor_set(x_466, 1, x_465); -return x_466; +lean_object* x_576; lean_object* x_577; lean_object* x_578; +x_576 = lean_ctor_get(x_527, 0); +x_577 = lean_ctor_get(x_527, 1); +lean_inc(x_577); +lean_inc(x_576); +lean_dec(x_527); +x_578 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_578, 0, x_576); +lean_ctor_set(x_578, 1, x_577); +return x_578; } } } } -case 7: +case 4: { -lean_object* x_480; lean_object* x_511; lean_object* x_545; uint8_t x_546; +lean_object* x_592; lean_object* x_671; lean_object* x_725; uint8_t x_726; lean_dec(x_4); -x_545 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_546 = l_Lean_Expr_isConstOf(x_2, x_545); -if (x_546 == 0) +x_725 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_726 = l_Lean_Expr_isConstOf(x_2, x_725); +if (x_726 == 0) { -lean_object* x_547; -x_547 = lean_box(0); -x_511 = x_547; -goto block_544; +lean_object* x_727; +x_727 = lean_box(0); +x_671 = x_727; +goto block_724; } else { -lean_object* x_548; lean_object* x_549; uint8_t x_550; -x_548 = lean_array_get_size(x_3); -x_549 = lean_unsigned_to_nat(2u); -x_550 = lean_nat_dec_eq(x_548, x_549); -if (x_550 == 0) +lean_object* x_728; lean_object* x_729; uint8_t x_730; +x_728 = lean_array_get_size(x_3); +x_729 = lean_unsigned_to_nat(2u); +x_730 = lean_nat_dec_eq(x_728, x_729); +if (x_730 == 0) { -lean_object* x_551; -lean_dec(x_548); -x_551 = lean_box(0); -x_511 = x_551; -goto block_544; +lean_object* x_731; +lean_dec(x_728); +x_731 = lean_box(0); +x_671 = x_731; +goto block_724; } else { -lean_object* x_552; uint8_t x_553; -x_552 = lean_unsigned_to_nat(0u); -x_553 = lean_nat_dec_lt(x_552, x_548); -lean_dec(x_548); -if (x_553 == 0) +lean_object* x_732; uint8_t x_733; +x_732 = lean_unsigned_to_nat(0u); +x_733 = lean_nat_dec_lt(x_732, x_728); +lean_dec(x_728); +if (x_733 == 0) { -lean_object* x_554; lean_object* x_555; -x_554 = l_Lean_instInhabitedExpr; -x_555 = l_outOfBounds___rarg(x_554); -x_480 = x_555; -goto block_510; +lean_object* x_734; lean_object* x_735; +x_734 = l_Lean_instInhabitedExpr; +x_735 = l_outOfBounds___rarg(x_734); +x_592 = x_735; +goto block_670; } else { -lean_object* x_556; -x_556 = lean_array_fget(x_3, x_552); -x_480 = x_556; -goto block_510; +lean_object* x_736; +x_736 = lean_array_fget(x_3, x_732); +x_592 = x_736; +goto block_670; } } } -block_510: +block_670: { -lean_object* x_481; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_480); -x_481 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_480, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_481) == 0) +lean_object* x_593; +lean_inc(x_592); +x_593 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_592, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_593) == 0) +{ +uint8_t x_594; +x_594 = !lean_is_exclusive(x_593); +if (x_594 == 0) +{ +lean_object* x_595; uint8_t x_596; +x_595 = lean_ctor_get(x_593, 0); +x_596 = !lean_is_exclusive(x_595); +if (x_596 == 0) +{ +lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; +x_597 = lean_ctor_get(x_595, 0); +x_598 = lean_ctor_get(x_595, 1); +x_599 = lean_ctor_get(x_598, 2); +lean_inc(x_599); +lean_inc(x_599); +x_600 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_599, x_597); +if (lean_obj_tag(x_600) == 0) { -lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; -x_482 = lean_ctor_get(x_481, 0); -lean_inc(x_482); -x_483 = lean_ctor_get(x_481, 1); -lean_inc(x_483); -lean_dec(x_481); -x_484 = lean_ctor_get(x_482, 0); -lean_inc(x_484); -x_485 = lean_ctor_get(x_482, 1); -lean_inc(x_485); -lean_dec(x_482); -x_486 = lean_ctor_get(x_485, 2); -lean_inc(x_486); -lean_inc(x_486); -x_487 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_486, x_484); -if (lean_obj_tag(x_487) == 0) -{ -size_t x_488; size_t x_489; uint8_t x_490; -x_488 = lean_ptr_addr(x_480); -lean_dec(x_480); -x_489 = lean_ptr_addr(x_484); -x_490 = lean_usize_dec_eq(x_488, x_489); -if (x_490 == 0) -{ -lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; -x_491 = lean_unsigned_to_nat(0u); -lean_inc(x_484); -x_492 = lean_array_set(x_3, x_491, x_484); -x_493 = l_Lean_mkAppN(x_2, x_492); -lean_dec(x_492); -x_494 = lean_ctor_get(x_485, 0); -lean_inc(x_494); -x_495 = lean_ctor_get(x_485, 1); -lean_inc(x_495); -lean_dec(x_485); -lean_inc(x_493); -x_496 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_486, x_484, x_493); -x_497 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_497, 0, x_494); -lean_ctor_set(x_497, 1, x_495); -lean_ctor_set(x_497, 2, x_496); -x_498 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_493, x_5, x_497, x_7, x_8, x_9, x_10, x_483); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_498; +size_t x_601; size_t x_602; uint8_t x_603; +x_601 = lean_ptr_addr(x_592); +lean_dec(x_592); +x_602 = lean_ptr_addr(x_597); +x_603 = lean_usize_dec_eq(x_601, x_602); +if (x_603 == 0) +{ +lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; +lean_dec(x_1); +x_604 = lean_unsigned_to_nat(0u); +lean_inc(x_597); +x_605 = lean_array_set(x_3, x_604, x_597); +x_606 = l_Lean_mkAppN(x_2, x_605); +lean_dec(x_605); +x_607 = lean_ctor_get(x_598, 0); +lean_inc(x_607); +x_608 = lean_ctor_get(x_598, 1); +lean_inc(x_608); +lean_dec(x_598); +lean_inc(x_606); +x_609 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_599, x_597, x_606); +x_610 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_610, 0, x_607); +lean_ctor_set(x_610, 1, x_608); +lean_ctor_set(x_610, 2, x_609); +lean_ctor_set(x_595, 1, x_610); +lean_ctor_set(x_595, 0, x_606); +return x_593; } else { -lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; +lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_dec(x_3); lean_dec(x_2); -x_499 = lean_ctor_get(x_485, 0); -lean_inc(x_499); -x_500 = lean_ctor_get(x_485, 1); -lean_inc(x_500); -lean_dec(x_485); +x_611 = lean_ctor_get(x_598, 0); +lean_inc(x_611); +x_612 = lean_ctor_get(x_598, 1); +lean_inc(x_612); +lean_dec(x_598); lean_inc(x_1); -x_501 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_486, x_484, x_1); -x_502 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_502, 0, x_499); -lean_ctor_set(x_502, 1, x_500); -lean_ctor_set(x_502, 2, x_501); -lean_inc(x_1); -x_503 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_502, x_7, x_8, x_9, x_10, x_483); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_503; +x_613 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_599, x_597, x_1); +x_614 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_614, 0, x_611); +lean_ctor_set(x_614, 1, x_612); +lean_ctor_set(x_614, 2, x_613); +lean_ctor_set(x_595, 1, x_614); +lean_ctor_set(x_595, 0, x_1); +return x_593; } } else { -lean_object* x_504; lean_object* x_505; -lean_dec(x_486); -lean_dec(x_484); -lean_dec(x_480); +lean_object* x_615; +lean_dec(x_599); +lean_dec(x_597); +lean_dec(x_592); lean_dec(x_3); lean_dec(x_2); -x_504 = lean_ctor_get(x_487, 0); -lean_inc(x_504); -lean_dec(x_487); -x_505 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_504, x_5, x_485, x_7, x_8, x_9, x_10, x_483); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_505; +lean_dec(x_1); +x_615 = lean_ctor_get(x_600, 0); +lean_inc(x_615); +lean_dec(x_600); +lean_ctor_set(x_595, 0, x_615); +return x_593; } } else { -uint8_t x_506; -lean_dec(x_480); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; +x_616 = lean_ctor_get(x_595, 0); +x_617 = lean_ctor_get(x_595, 1); +lean_inc(x_617); +lean_inc(x_616); +lean_dec(x_595); +x_618 = lean_ctor_get(x_617, 2); +lean_inc(x_618); +lean_inc(x_618); +x_619 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_618, x_616); +if (lean_obj_tag(x_619) == 0) +{ +size_t x_620; size_t x_621; uint8_t x_622; +x_620 = lean_ptr_addr(x_592); +lean_dec(x_592); +x_621 = lean_ptr_addr(x_616); +x_622 = lean_usize_dec_eq(x_620, x_621); +if (x_622 == 0) +{ +lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; +lean_dec(x_1); +x_623 = lean_unsigned_to_nat(0u); +lean_inc(x_616); +x_624 = lean_array_set(x_3, x_623, x_616); +x_625 = l_Lean_mkAppN(x_2, x_624); +lean_dec(x_624); +x_626 = lean_ctor_get(x_617, 0); +lean_inc(x_626); +x_627 = lean_ctor_get(x_617, 1); +lean_inc(x_627); +lean_dec(x_617); +lean_inc(x_625); +x_628 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_618, x_616, x_625); +x_629 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_629, 0, x_626); +lean_ctor_set(x_629, 1, x_627); +lean_ctor_set(x_629, 2, x_628); +x_630 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_630, 0, x_625); +lean_ctor_set(x_630, 1, x_629); +lean_ctor_set(x_593, 0, x_630); +return x_593; +} +else +{ +lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_506 = !lean_is_exclusive(x_481); -if (x_506 == 0) -{ -return x_481; +x_631 = lean_ctor_get(x_617, 0); +lean_inc(x_631); +x_632 = lean_ctor_get(x_617, 1); +lean_inc(x_632); +lean_dec(x_617); +lean_inc(x_1); +x_633 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_618, x_616, x_1); +x_634 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_634, 0, x_631); +lean_ctor_set(x_634, 1, x_632); +lean_ctor_set(x_634, 2, x_633); +x_635 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_635, 0, x_1); +lean_ctor_set(x_635, 1, x_634); +lean_ctor_set(x_593, 0, x_635); +return x_593; +} } else { -lean_object* x_507; lean_object* x_508; lean_object* x_509; -x_507 = lean_ctor_get(x_481, 0); -x_508 = lean_ctor_get(x_481, 1); -lean_inc(x_508); -lean_inc(x_507); -lean_dec(x_481); -x_509 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_509, 0, x_507); -lean_ctor_set(x_509, 1, x_508); -return x_509; -} +lean_object* x_636; lean_object* x_637; +lean_dec(x_618); +lean_dec(x_616); +lean_dec(x_592); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_636 = lean_ctor_get(x_619, 0); +lean_inc(x_636); +lean_dec(x_619); +x_637 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_637, 0, x_636); +lean_ctor_set(x_637, 1, x_617); +lean_ctor_set(x_593, 0, x_637); +return x_593; } } -block_544: -{ -lean_object* x_512; -lean_dec(x_511); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_512 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_512) == 0) -{ -lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; uint8_t x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; -x_513 = lean_ctor_get(x_512, 0); -lean_inc(x_513); -x_514 = lean_ctor_get(x_512, 1); -lean_inc(x_514); -lean_dec(x_512); -x_515 = lean_ctor_get(x_513, 0); -lean_inc(x_515); -lean_dec(x_513); -x_516 = lean_array_get_size(x_3); -x_517 = lean_unsigned_to_nat(0u); -x_518 = lean_unsigned_to_nat(1u); -x_519 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_519, 0, x_517); -lean_ctor_set(x_519, 1, x_516); -lean_ctor_set(x_519, 2, x_518); -x_520 = 0; -x_521 = lean_box(x_520); -x_522 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_522, 0, x_3); -lean_ctor_set(x_522, 1, x_521); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_523 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_515, x_519, x_519, x_522, x_517, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_514); -lean_dec(x_519); -lean_dec(x_515); -if (lean_obj_tag(x_523) == 0) -{ -lean_object* x_524; lean_object* x_525; lean_object* x_526; uint8_t x_527; -x_524 = lean_ctor_get(x_523, 0); -lean_inc(x_524); -x_525 = lean_ctor_get(x_524, 0); -lean_inc(x_525); -x_526 = lean_ctor_get(x_525, 1); -lean_inc(x_526); -x_527 = lean_unbox(x_526); -lean_dec(x_526); -if (x_527 == 0) -{ -lean_object* x_528; lean_object* x_529; lean_object* x_530; -lean_dec(x_525); -lean_dec(x_2); -x_528 = lean_ctor_get(x_523, 1); -lean_inc(x_528); -lean_dec(x_523); -x_529 = lean_ctor_get(x_524, 1); -lean_inc(x_529); -lean_dec(x_524); -lean_inc(x_1); -x_530 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_529, x_7, x_8, x_9, x_10, x_528); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_530; } else { -lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; -x_531 = lean_ctor_get(x_523, 1); -lean_inc(x_531); -lean_dec(x_523); -x_532 = lean_ctor_get(x_524, 1); -lean_inc(x_532); -lean_dec(x_524); -x_533 = lean_ctor_get(x_525, 0); -lean_inc(x_533); -lean_dec(x_525); -x_534 = l_Lean_mkAppN(x_2, x_533); -lean_dec(x_533); -x_535 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_534, x_5, x_532, x_7, x_8, x_9, x_10, x_531); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_535; -} +lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; +x_638 = lean_ctor_get(x_593, 0); +x_639 = lean_ctor_get(x_593, 1); +lean_inc(x_639); +lean_inc(x_638); +lean_dec(x_593); +x_640 = lean_ctor_get(x_638, 0); +lean_inc(x_640); +x_641 = lean_ctor_get(x_638, 1); +lean_inc(x_641); +if (lean_is_exclusive(x_638)) { + lean_ctor_release(x_638, 0); + lean_ctor_release(x_638, 1); + x_642 = x_638; +} else { + lean_dec_ref(x_638); + x_642 = lean_box(0); } -else +x_643 = lean_ctor_get(x_641, 2); +lean_inc(x_643); +lean_inc(x_643); +x_644 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_643, x_640); +if (lean_obj_tag(x_644) == 0) { -uint8_t x_536; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_536 = !lean_is_exclusive(x_523); -if (x_536 == 0) +size_t x_645; size_t x_646; uint8_t x_647; +x_645 = lean_ptr_addr(x_592); +lean_dec(x_592); +x_646 = lean_ptr_addr(x_640); +x_647 = lean_usize_dec_eq(x_645, x_646); +if (x_647 == 0) { -return x_523; +lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; +lean_dec(x_1); +x_648 = lean_unsigned_to_nat(0u); +lean_inc(x_640); +x_649 = lean_array_set(x_3, x_648, x_640); +x_650 = l_Lean_mkAppN(x_2, x_649); +lean_dec(x_649); +x_651 = lean_ctor_get(x_641, 0); +lean_inc(x_651); +x_652 = lean_ctor_get(x_641, 1); +lean_inc(x_652); +lean_dec(x_641); +lean_inc(x_650); +x_653 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_643, x_640, x_650); +x_654 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_654, 0, x_651); +lean_ctor_set(x_654, 1, x_652); +lean_ctor_set(x_654, 2, x_653); +if (lean_is_scalar(x_642)) { + x_655 = lean_alloc_ctor(0, 2, 0); +} else { + x_655 = x_642; +} +lean_ctor_set(x_655, 0, x_650); +lean_ctor_set(x_655, 1, x_654); +x_656 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_656, 0, x_655); +lean_ctor_set(x_656, 1, x_639); +return x_656; } else { -lean_object* x_537; lean_object* x_538; lean_object* x_539; -x_537 = lean_ctor_get(x_523, 0); -x_538 = lean_ctor_get(x_523, 1); -lean_inc(x_538); -lean_inc(x_537); -lean_dec(x_523); -x_539 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_539, 0, x_537); -lean_ctor_set(x_539, 1, x_538); -return x_539; +lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; +lean_dec(x_3); +lean_dec(x_2); +x_657 = lean_ctor_get(x_641, 0); +lean_inc(x_657); +x_658 = lean_ctor_get(x_641, 1); +lean_inc(x_658); +lean_dec(x_641); +lean_inc(x_1); +x_659 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_643, x_640, x_1); +x_660 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_660, 0, x_657); +lean_ctor_set(x_660, 1, x_658); +lean_ctor_set(x_660, 2, x_659); +if (lean_is_scalar(x_642)) { + x_661 = lean_alloc_ctor(0, 2, 0); +} else { + x_661 = x_642; } +lean_ctor_set(x_661, 0, x_1); +lean_ctor_set(x_661, 1, x_660); +x_662 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_662, 0, x_661); +lean_ctor_set(x_662, 1, x_639); +return x_662; } } else { -uint8_t x_540; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_663; lean_object* x_664; lean_object* x_665; +lean_dec(x_643); +lean_dec(x_640); +lean_dec(x_592); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_540 = !lean_is_exclusive(x_512); -if (x_540 == 0) -{ -return x_512; -} -else -{ -lean_object* x_541; lean_object* x_542; lean_object* x_543; -x_541 = lean_ctor_get(x_512, 0); -x_542 = lean_ctor_get(x_512, 1); -lean_inc(x_542); -lean_inc(x_541); -lean_dec(x_512); -x_543 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_543, 0, x_541); -lean_ctor_set(x_543, 1, x_542); -return x_543; -} -} +x_663 = lean_ctor_get(x_644, 0); +lean_inc(x_663); +lean_dec(x_644); +if (lean_is_scalar(x_642)) { + x_664 = lean_alloc_ctor(0, 2, 0); +} else { + x_664 = x_642; } +lean_ctor_set(x_664, 0, x_663); +lean_ctor_set(x_664, 1, x_641); +x_665 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_665, 0, x_664); +lean_ctor_set(x_665, 1, x_639); +return x_665; } -case 8: -{ -lean_object* x_557; lean_object* x_588; lean_object* x_622; uint8_t x_623; -lean_dec(x_4); -x_622 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_623 = l_Lean_Expr_isConstOf(x_2, x_622); -if (x_623 == 0) -{ -lean_object* x_624; -x_624 = lean_box(0); -x_588 = x_624; -goto block_621; } -else -{ -lean_object* x_625; lean_object* x_626; uint8_t x_627; -x_625 = lean_array_get_size(x_3); -x_626 = lean_unsigned_to_nat(2u); -x_627 = lean_nat_dec_eq(x_625, x_626); -if (x_627 == 0) -{ -lean_object* x_628; -lean_dec(x_625); -x_628 = lean_box(0); -x_588 = x_628; -goto block_621; } else { -lean_object* x_629; uint8_t x_630; -x_629 = lean_unsigned_to_nat(0u); -x_630 = lean_nat_dec_lt(x_629, x_625); -lean_dec(x_625); -if (x_630 == 0) +uint8_t x_666; +lean_dec(x_592); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_666 = !lean_is_exclusive(x_593); +if (x_666 == 0) { -lean_object* x_631; lean_object* x_632; -x_631 = l_Lean_instInhabitedExpr; -x_632 = l_outOfBounds___rarg(x_631); -x_557 = x_632; -goto block_587; +return x_593; } else { -lean_object* x_633; -x_633 = lean_array_fget(x_3, x_629); -x_557 = x_633; -goto block_587; +lean_object* x_667; lean_object* x_668; lean_object* x_669; +x_667 = lean_ctor_get(x_593, 0); +x_668 = lean_ctor_get(x_593, 1); +lean_inc(x_668); +lean_inc(x_667); +lean_dec(x_593); +x_669 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_669, 0, x_667); +lean_ctor_set(x_669, 1, x_668); +return x_669; } } } -block_587: +block_724: { -lean_object* x_558; +lean_object* x_672; +lean_dec(x_671); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_557); -x_558 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_557, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_558) == 0) -{ -lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; -x_559 = lean_ctor_get(x_558, 0); -lean_inc(x_559); -x_560 = lean_ctor_get(x_558, 1); -lean_inc(x_560); -lean_dec(x_558); -x_561 = lean_ctor_get(x_559, 0); -lean_inc(x_561); -x_562 = lean_ctor_get(x_559, 1); -lean_inc(x_562); -lean_dec(x_559); -x_563 = lean_ctor_get(x_562, 2); -lean_inc(x_563); -lean_inc(x_563); -x_564 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_563, x_561); -if (lean_obj_tag(x_564) == 0) -{ -size_t x_565; size_t x_566; uint8_t x_567; -x_565 = lean_ptr_addr(x_557); -lean_dec(x_557); -x_566 = lean_ptr_addr(x_561); -x_567 = lean_usize_dec_eq(x_565, x_566); -if (x_567 == 0) +lean_inc(x_2); +x_672 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_672) == 0) +{ +lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; uint8_t x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; +x_673 = lean_ctor_get(x_672, 0); +lean_inc(x_673); +x_674 = lean_ctor_get(x_672, 1); +lean_inc(x_674); +lean_dec(x_672); +x_675 = lean_ctor_get(x_673, 0); +lean_inc(x_675); +lean_dec(x_673); +x_676 = lean_array_get_size(x_3); +x_677 = lean_unsigned_to_nat(0u); +x_678 = lean_unsigned_to_nat(1u); +lean_inc(x_676); +x_679 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_679, 0, x_677); +lean_ctor_set(x_679, 1, x_676); +lean_ctor_set(x_679, 2, x_678); +x_680 = 0; +x_681 = lean_box(x_680); +lean_inc(x_3); +x_682 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_682, 0, x_3); +lean_ctor_set(x_682, 1, x_681); +lean_inc(x_2); +x_683 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_675, x_676, x_679, x_679, x_682, x_677, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_674); +lean_dec(x_679); +lean_dec(x_676); +lean_dec(x_675); +lean_dec(x_3); +if (lean_obj_tag(x_683) == 0) { -lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; -x_568 = lean_unsigned_to_nat(0u); -lean_inc(x_561); -x_569 = lean_array_set(x_3, x_568, x_561); -x_570 = l_Lean_mkAppN(x_2, x_569); -lean_dec(x_569); -x_571 = lean_ctor_get(x_562, 0); -lean_inc(x_571); -x_572 = lean_ctor_get(x_562, 1); -lean_inc(x_572); -lean_dec(x_562); -lean_inc(x_570); -x_573 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_563, x_561, x_570); -x_574 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_574, 0, x_571); -lean_ctor_set(x_574, 1, x_572); -lean_ctor_set(x_574, 2, x_573); -x_575 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_570, x_5, x_574, x_7, x_8, x_9, x_10, x_560); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_575; +lean_object* x_684; lean_object* x_685; lean_object* x_686; uint8_t x_687; +x_684 = lean_ctor_get(x_683, 0); +lean_inc(x_684); +x_685 = lean_ctor_get(x_684, 0); +lean_inc(x_685); +x_686 = lean_ctor_get(x_685, 1); +lean_inc(x_686); +x_687 = lean_unbox(x_686); +lean_dec(x_686); +if (x_687 == 0) +{ +uint8_t x_688; +lean_dec(x_685); +lean_dec(x_2); +x_688 = !lean_is_exclusive(x_683); +if (x_688 == 0) +{ +lean_object* x_689; uint8_t x_690; +x_689 = lean_ctor_get(x_683, 0); +lean_dec(x_689); +x_690 = !lean_is_exclusive(x_684); +if (x_690 == 0) +{ +lean_object* x_691; +x_691 = lean_ctor_get(x_684, 0); +lean_dec(x_691); +lean_ctor_set(x_684, 0, x_1); +return x_683; } else { -lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; -lean_dec(x_3); -lean_dec(x_2); -x_576 = lean_ctor_get(x_562, 0); -lean_inc(x_576); -x_577 = lean_ctor_get(x_562, 1); -lean_inc(x_577); -lean_dec(x_562); -lean_inc(x_1); -x_578 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_563, x_561, x_1); -x_579 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_579, 0, x_576); -lean_ctor_set(x_579, 1, x_577); -lean_ctor_set(x_579, 2, x_578); -lean_inc(x_1); -x_580 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_579, x_7, x_8, x_9, x_10, x_560); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_580; +lean_object* x_692; lean_object* x_693; +x_692 = lean_ctor_get(x_684, 1); +lean_inc(x_692); +lean_dec(x_684); +x_693 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_693, 0, x_1); +lean_ctor_set(x_693, 1, x_692); +lean_ctor_set(x_683, 0, x_693); +return x_683; } } else { -lean_object* x_581; lean_object* x_582; -lean_dec(x_563); -lean_dec(x_561); -lean_dec(x_557); -lean_dec(x_3); -lean_dec(x_2); -x_581 = lean_ctor_get(x_564, 0); -lean_inc(x_581); -lean_dec(x_564); -x_582 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_581, x_5, x_562, x_7, x_8, x_9, x_10, x_560); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_582; +lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; +x_694 = lean_ctor_get(x_683, 1); +lean_inc(x_694); +lean_dec(x_683); +x_695 = lean_ctor_get(x_684, 1); +lean_inc(x_695); +if (lean_is_exclusive(x_684)) { + lean_ctor_release(x_684, 0); + lean_ctor_release(x_684, 1); + x_696 = x_684; +} else { + lean_dec_ref(x_684); + x_696 = lean_box(0); +} +if (lean_is_scalar(x_696)) { + x_697 = lean_alloc_ctor(0, 2, 0); +} else { + x_697 = x_696; +} +lean_ctor_set(x_697, 0, x_1); +lean_ctor_set(x_697, 1, x_695); +x_698 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_698, 0, x_697); +lean_ctor_set(x_698, 1, x_694); +return x_698; } } else { -uint8_t x_583; -lean_dec(x_557); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); +uint8_t x_699; lean_dec(x_1); -x_583 = !lean_is_exclusive(x_558); -if (x_583 == 0) +x_699 = !lean_is_exclusive(x_683); +if (x_699 == 0) { -return x_558; -} -else +lean_object* x_700; uint8_t x_701; +x_700 = lean_ctor_get(x_683, 0); +lean_dec(x_700); +x_701 = !lean_is_exclusive(x_684); +if (x_701 == 0) { -lean_object* x_584; lean_object* x_585; lean_object* x_586; -x_584 = lean_ctor_get(x_558, 0); -x_585 = lean_ctor_get(x_558, 1); -lean_inc(x_585); -lean_inc(x_584); -lean_dec(x_558); -x_586 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_586, 0, x_584); -lean_ctor_set(x_586, 1, x_585); -return x_586; -} +lean_object* x_702; lean_object* x_703; lean_object* x_704; +x_702 = lean_ctor_get(x_684, 0); +lean_dec(x_702); +x_703 = lean_ctor_get(x_685, 0); +lean_inc(x_703); +lean_dec(x_685); +x_704 = l_Lean_mkAppN(x_2, x_703); +lean_dec(x_703); +lean_ctor_set(x_684, 0, x_704); +return x_683; +} +else +{ +lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; +x_705 = lean_ctor_get(x_684, 1); +lean_inc(x_705); +lean_dec(x_684); +x_706 = lean_ctor_get(x_685, 0); +lean_inc(x_706); +lean_dec(x_685); +x_707 = l_Lean_mkAppN(x_2, x_706); +lean_dec(x_706); +x_708 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_708, 0, x_707); +lean_ctor_set(x_708, 1, x_705); +lean_ctor_set(x_683, 0, x_708); +return x_683; +} +} +else +{ +lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; +x_709 = lean_ctor_get(x_683, 1); +lean_inc(x_709); +lean_dec(x_683); +x_710 = lean_ctor_get(x_684, 1); +lean_inc(x_710); +if (lean_is_exclusive(x_684)) { + lean_ctor_release(x_684, 0); + lean_ctor_release(x_684, 1); + x_711 = x_684; +} else { + lean_dec_ref(x_684); + x_711 = lean_box(0); } +x_712 = lean_ctor_get(x_685, 0); +lean_inc(x_712); +lean_dec(x_685); +x_713 = l_Lean_mkAppN(x_2, x_712); +lean_dec(x_712); +if (lean_is_scalar(x_711)) { + x_714 = lean_alloc_ctor(0, 2, 0); +} else { + x_714 = x_711; } -block_621: -{ -lean_object* x_589; -lean_dec(x_588); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_589 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_589) == 0) -{ -lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; uint8_t x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; -x_590 = lean_ctor_get(x_589, 0); -lean_inc(x_590); -x_591 = lean_ctor_get(x_589, 1); -lean_inc(x_591); -lean_dec(x_589); -x_592 = lean_ctor_get(x_590, 0); -lean_inc(x_592); -lean_dec(x_590); -x_593 = lean_array_get_size(x_3); -x_594 = lean_unsigned_to_nat(0u); -x_595 = lean_unsigned_to_nat(1u); -x_596 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_596, 0, x_594); -lean_ctor_set(x_596, 1, x_593); -lean_ctor_set(x_596, 2, x_595); -x_597 = 0; -x_598 = lean_box(x_597); -x_599 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_599, 0, x_3); -lean_ctor_set(x_599, 1, x_598); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_600 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_592, x_596, x_596, x_599, x_594, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_591); -lean_dec(x_596); -lean_dec(x_592); -if (lean_obj_tag(x_600) == 0) -{ -lean_object* x_601; lean_object* x_602; lean_object* x_603; uint8_t x_604; -x_601 = lean_ctor_get(x_600, 0); -lean_inc(x_601); -x_602 = lean_ctor_get(x_601, 0); -lean_inc(x_602); -x_603 = lean_ctor_get(x_602, 1); -lean_inc(x_603); -x_604 = lean_unbox(x_603); -lean_dec(x_603); -if (x_604 == 0) -{ -lean_object* x_605; lean_object* x_606; lean_object* x_607; -lean_dec(x_602); -lean_dec(x_2); -x_605 = lean_ctor_get(x_600, 1); -lean_inc(x_605); -lean_dec(x_600); -x_606 = lean_ctor_get(x_601, 1); -lean_inc(x_606); -lean_dec(x_601); -lean_inc(x_1); -x_607 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_606, x_7, x_8, x_9, x_10, x_605); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_607; +lean_ctor_set(x_714, 0, x_713); +lean_ctor_set(x_714, 1, x_710); +x_715 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_715, 0, x_714); +lean_ctor_set(x_715, 1, x_709); +return x_715; } -else -{ -lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; -x_608 = lean_ctor_get(x_600, 1); -lean_inc(x_608); -lean_dec(x_600); -x_609 = lean_ctor_get(x_601, 1); -lean_inc(x_609); -lean_dec(x_601); -x_610 = lean_ctor_get(x_602, 0); -lean_inc(x_610); -lean_dec(x_602); -x_611 = l_Lean_mkAppN(x_2, x_610); -lean_dec(x_610); -x_612 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_611, x_5, x_609, x_7, x_8, x_9, x_10, x_608); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_612; } } else { -uint8_t x_613; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +uint8_t x_716; lean_dec(x_2); lean_dec(x_1); -x_613 = !lean_is_exclusive(x_600); -if (x_613 == 0) +x_716 = !lean_is_exclusive(x_683); +if (x_716 == 0) { -return x_600; +return x_683; } else { -lean_object* x_614; lean_object* x_615; lean_object* x_616; -x_614 = lean_ctor_get(x_600, 0); -x_615 = lean_ctor_get(x_600, 1); -lean_inc(x_615); -lean_inc(x_614); -lean_dec(x_600); -x_616 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_616, 0, x_614); -lean_ctor_set(x_616, 1, x_615); -return x_616; +lean_object* x_717; lean_object* x_718; lean_object* x_719; +x_717 = lean_ctor_get(x_683, 0); +x_718 = lean_ctor_get(x_683, 1); +lean_inc(x_718); +lean_inc(x_717); +lean_dec(x_683); +x_719 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_719, 0, x_717); +lean_ctor_set(x_719, 1, x_718); +return x_719; } } } else { -uint8_t x_617; +uint8_t x_720; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -9619,1400 +9453,4903 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_617 = !lean_is_exclusive(x_589); -if (x_617 == 0) +x_720 = !lean_is_exclusive(x_672); +if (x_720 == 0) { -return x_589; +return x_672; } else { -lean_object* x_618; lean_object* x_619; lean_object* x_620; -x_618 = lean_ctor_get(x_589, 0); -x_619 = lean_ctor_get(x_589, 1); -lean_inc(x_619); -lean_inc(x_618); -lean_dec(x_589); -x_620 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_620, 0, x_618); -lean_ctor_set(x_620, 1, x_619); -return x_620; +lean_object* x_721; lean_object* x_722; lean_object* x_723; +x_721 = lean_ctor_get(x_672, 0); +x_722 = lean_ctor_get(x_672, 1); +lean_inc(x_722); +lean_inc(x_721); +lean_dec(x_672); +x_723 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_723, 0, x_721); +lean_ctor_set(x_723, 1, x_722); +return x_723; } } } } -case 9: +case 5: { -lean_object* x_634; lean_object* x_665; lean_object* x_699; uint8_t x_700; +lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; +x_737 = lean_ctor_get(x_2, 0); +lean_inc(x_737); +x_738 = lean_ctor_get(x_2, 1); +lean_inc(x_738); +lean_dec(x_2); +x_739 = lean_array_set(x_3, x_4, x_738); +x_740 = lean_unsigned_to_nat(1u); +x_741 = lean_nat_sub(x_4, x_740); +lean_dec(x_4); +x_2 = x_737; +x_3 = x_739; +x_4 = x_741; +goto _start; +} +case 6: +{ +lean_object* x_743; lean_object* x_822; lean_object* x_876; uint8_t x_877; lean_dec(x_4); -x_699 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_700 = l_Lean_Expr_isConstOf(x_2, x_699); -if (x_700 == 0) +x_876 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_877 = l_Lean_Expr_isConstOf(x_2, x_876); +if (x_877 == 0) { -lean_object* x_701; -x_701 = lean_box(0); -x_665 = x_701; -goto block_698; +lean_object* x_878; +x_878 = lean_box(0); +x_822 = x_878; +goto block_875; } else { -lean_object* x_702; lean_object* x_703; uint8_t x_704; -x_702 = lean_array_get_size(x_3); -x_703 = lean_unsigned_to_nat(2u); -x_704 = lean_nat_dec_eq(x_702, x_703); -if (x_704 == 0) +lean_object* x_879; lean_object* x_880; uint8_t x_881; +x_879 = lean_array_get_size(x_3); +x_880 = lean_unsigned_to_nat(2u); +x_881 = lean_nat_dec_eq(x_879, x_880); +if (x_881 == 0) { -lean_object* x_705; -lean_dec(x_702); -x_705 = lean_box(0); -x_665 = x_705; -goto block_698; +lean_object* x_882; +lean_dec(x_879); +x_882 = lean_box(0); +x_822 = x_882; +goto block_875; } else { -lean_object* x_706; uint8_t x_707; -x_706 = lean_unsigned_to_nat(0u); -x_707 = lean_nat_dec_lt(x_706, x_702); -lean_dec(x_702); -if (x_707 == 0) +lean_object* x_883; uint8_t x_884; +x_883 = lean_unsigned_to_nat(0u); +x_884 = lean_nat_dec_lt(x_883, x_879); +lean_dec(x_879); +if (x_884 == 0) { -lean_object* x_708; lean_object* x_709; -x_708 = l_Lean_instInhabitedExpr; -x_709 = l_outOfBounds___rarg(x_708); -x_634 = x_709; -goto block_664; +lean_object* x_885; lean_object* x_886; +x_885 = l_Lean_instInhabitedExpr; +x_886 = l_outOfBounds___rarg(x_885); +x_743 = x_886; +goto block_821; } else { -lean_object* x_710; -x_710 = lean_array_fget(x_3, x_706); -x_634 = x_710; -goto block_664; +lean_object* x_887; +x_887 = lean_array_fget(x_3, x_883); +x_743 = x_887; +goto block_821; } } } -block_664: +block_821: { -lean_object* x_635; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_634); -x_635 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_634, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_635) == 0) +lean_object* x_744; +lean_inc(x_743); +x_744 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_743, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_744) == 0) { -lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; -x_636 = lean_ctor_get(x_635, 0); -lean_inc(x_636); -x_637 = lean_ctor_get(x_635, 1); -lean_inc(x_637); -lean_dec(x_635); -x_638 = lean_ctor_get(x_636, 0); -lean_inc(x_638); -x_639 = lean_ctor_get(x_636, 1); -lean_inc(x_639); -lean_dec(x_636); -x_640 = lean_ctor_get(x_639, 2); -lean_inc(x_640); -lean_inc(x_640); -x_641 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_640, x_638); -if (lean_obj_tag(x_641) == 0) -{ -size_t x_642; size_t x_643; uint8_t x_644; -x_642 = lean_ptr_addr(x_634); -lean_dec(x_634); -x_643 = lean_ptr_addr(x_638); -x_644 = lean_usize_dec_eq(x_642, x_643); -if (x_644 == 0) -{ -lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; -x_645 = lean_unsigned_to_nat(0u); -lean_inc(x_638); -x_646 = lean_array_set(x_3, x_645, x_638); -x_647 = l_Lean_mkAppN(x_2, x_646); -lean_dec(x_646); -x_648 = lean_ctor_get(x_639, 0); -lean_inc(x_648); -x_649 = lean_ctor_get(x_639, 1); -lean_inc(x_649); -lean_dec(x_639); -lean_inc(x_647); -x_650 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_640, x_638, x_647); -x_651 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_651, 0, x_648); -lean_ctor_set(x_651, 1, x_649); -lean_ctor_set(x_651, 2, x_650); -x_652 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_647, x_5, x_651, x_7, x_8, x_9, x_10, x_637); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_652; +uint8_t x_745; +x_745 = !lean_is_exclusive(x_744); +if (x_745 == 0) +{ +lean_object* x_746; uint8_t x_747; +x_746 = lean_ctor_get(x_744, 0); +x_747 = !lean_is_exclusive(x_746); +if (x_747 == 0) +{ +lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; +x_748 = lean_ctor_get(x_746, 0); +x_749 = lean_ctor_get(x_746, 1); +x_750 = lean_ctor_get(x_749, 2); +lean_inc(x_750); +lean_inc(x_750); +x_751 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_750, x_748); +if (lean_obj_tag(x_751) == 0) +{ +size_t x_752; size_t x_753; uint8_t x_754; +x_752 = lean_ptr_addr(x_743); +lean_dec(x_743); +x_753 = lean_ptr_addr(x_748); +x_754 = lean_usize_dec_eq(x_752, x_753); +if (x_754 == 0) +{ +lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; +lean_dec(x_1); +x_755 = lean_unsigned_to_nat(0u); +lean_inc(x_748); +x_756 = lean_array_set(x_3, x_755, x_748); +x_757 = l_Lean_mkAppN(x_2, x_756); +lean_dec(x_756); +x_758 = lean_ctor_get(x_749, 0); +lean_inc(x_758); +x_759 = lean_ctor_get(x_749, 1); +lean_inc(x_759); +lean_dec(x_749); +lean_inc(x_757); +x_760 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_750, x_748, x_757); +x_761 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_761, 0, x_758); +lean_ctor_set(x_761, 1, x_759); +lean_ctor_set(x_761, 2, x_760); +lean_ctor_set(x_746, 1, x_761); +lean_ctor_set(x_746, 0, x_757); +return x_744; } else { -lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; +lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_dec(x_3); lean_dec(x_2); -x_653 = lean_ctor_get(x_639, 0); -lean_inc(x_653); -x_654 = lean_ctor_get(x_639, 1); -lean_inc(x_654); -lean_dec(x_639); -lean_inc(x_1); -x_655 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_640, x_638, x_1); -x_656 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_656, 0, x_653); -lean_ctor_set(x_656, 1, x_654); -lean_ctor_set(x_656, 2, x_655); +x_762 = lean_ctor_get(x_749, 0); +lean_inc(x_762); +x_763 = lean_ctor_get(x_749, 1); +lean_inc(x_763); +lean_dec(x_749); lean_inc(x_1); -x_657 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_656, x_7, x_8, x_9, x_10, x_637); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_657; +x_764 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_750, x_748, x_1); +x_765 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_765, 0, x_762); +lean_ctor_set(x_765, 1, x_763); +lean_ctor_set(x_765, 2, x_764); +lean_ctor_set(x_746, 1, x_765); +lean_ctor_set(x_746, 0, x_1); +return x_744; } } else { -lean_object* x_658; lean_object* x_659; -lean_dec(x_640); -lean_dec(x_638); -lean_dec(x_634); +lean_object* x_766; +lean_dec(x_750); +lean_dec(x_748); +lean_dec(x_743); lean_dec(x_3); lean_dec(x_2); -x_658 = lean_ctor_get(x_641, 0); -lean_inc(x_658); -lean_dec(x_641); -x_659 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_658, x_5, x_639, x_7, x_8, x_9, x_10, x_637); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_659; +lean_dec(x_1); +x_766 = lean_ctor_get(x_751, 0); +lean_inc(x_766); +lean_dec(x_751); +lean_ctor_set(x_746, 0, x_766); +return x_744; } } else { -uint8_t x_660; -lean_dec(x_634); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; +x_767 = lean_ctor_get(x_746, 0); +x_768 = lean_ctor_get(x_746, 1); +lean_inc(x_768); +lean_inc(x_767); +lean_dec(x_746); +x_769 = lean_ctor_get(x_768, 2); +lean_inc(x_769); +lean_inc(x_769); +x_770 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_769, x_767); +if (lean_obj_tag(x_770) == 0) +{ +size_t x_771; size_t x_772; uint8_t x_773; +x_771 = lean_ptr_addr(x_743); +lean_dec(x_743); +x_772 = lean_ptr_addr(x_767); +x_773 = lean_usize_dec_eq(x_771, x_772); +if (x_773 == 0) +{ +lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; +lean_dec(x_1); +x_774 = lean_unsigned_to_nat(0u); +lean_inc(x_767); +x_775 = lean_array_set(x_3, x_774, x_767); +x_776 = l_Lean_mkAppN(x_2, x_775); +lean_dec(x_775); +x_777 = lean_ctor_get(x_768, 0); +lean_inc(x_777); +x_778 = lean_ctor_get(x_768, 1); +lean_inc(x_778); +lean_dec(x_768); +lean_inc(x_776); +x_779 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_769, x_767, x_776); +x_780 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_780, 0, x_777); +lean_ctor_set(x_780, 1, x_778); +lean_ctor_set(x_780, 2, x_779); +x_781 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_781, 0, x_776); +lean_ctor_set(x_781, 1, x_780); +lean_ctor_set(x_744, 0, x_781); +return x_744; +} +else +{ +lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; +lean_dec(x_3); +lean_dec(x_2); +x_782 = lean_ctor_get(x_768, 0); +lean_inc(x_782); +x_783 = lean_ctor_get(x_768, 1); +lean_inc(x_783); +lean_dec(x_768); +lean_inc(x_1); +x_784 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_769, x_767, x_1); +x_785 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_785, 0, x_782); +lean_ctor_set(x_785, 1, x_783); +lean_ctor_set(x_785, 2, x_784); +x_786 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_786, 0, x_1); +lean_ctor_set(x_786, 1, x_785); +lean_ctor_set(x_744, 0, x_786); +return x_744; +} +} +else +{ +lean_object* x_787; lean_object* x_788; +lean_dec(x_769); +lean_dec(x_767); +lean_dec(x_743); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_787 = lean_ctor_get(x_770, 0); +lean_inc(x_787); +lean_dec(x_770); +x_788 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_788, 0, x_787); +lean_ctor_set(x_788, 1, x_768); +lean_ctor_set(x_744, 0, x_788); +return x_744; +} +} +} +else +{ +lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; +x_789 = lean_ctor_get(x_744, 0); +x_790 = lean_ctor_get(x_744, 1); +lean_inc(x_790); +lean_inc(x_789); +lean_dec(x_744); +x_791 = lean_ctor_get(x_789, 0); +lean_inc(x_791); +x_792 = lean_ctor_get(x_789, 1); +lean_inc(x_792); +if (lean_is_exclusive(x_789)) { + lean_ctor_release(x_789, 0); + lean_ctor_release(x_789, 1); + x_793 = x_789; +} else { + lean_dec_ref(x_789); + x_793 = lean_box(0); +} +x_794 = lean_ctor_get(x_792, 2); +lean_inc(x_794); +lean_inc(x_794); +x_795 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_794, x_791); +if (lean_obj_tag(x_795) == 0) +{ +size_t x_796; size_t x_797; uint8_t x_798; +x_796 = lean_ptr_addr(x_743); +lean_dec(x_743); +x_797 = lean_ptr_addr(x_791); +x_798 = lean_usize_dec_eq(x_796, x_797); +if (x_798 == 0) +{ +lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; +lean_dec(x_1); +x_799 = lean_unsigned_to_nat(0u); +lean_inc(x_791); +x_800 = lean_array_set(x_3, x_799, x_791); +x_801 = l_Lean_mkAppN(x_2, x_800); +lean_dec(x_800); +x_802 = lean_ctor_get(x_792, 0); +lean_inc(x_802); +x_803 = lean_ctor_get(x_792, 1); +lean_inc(x_803); +lean_dec(x_792); +lean_inc(x_801); +x_804 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_794, x_791, x_801); +x_805 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_805, 0, x_802); +lean_ctor_set(x_805, 1, x_803); +lean_ctor_set(x_805, 2, x_804); +if (lean_is_scalar(x_793)) { + x_806 = lean_alloc_ctor(0, 2, 0); +} else { + x_806 = x_793; +} +lean_ctor_set(x_806, 0, x_801); +lean_ctor_set(x_806, 1, x_805); +x_807 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_807, 0, x_806); +lean_ctor_set(x_807, 1, x_790); +return x_807; +} +else +{ +lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; +lean_dec(x_3); +lean_dec(x_2); +x_808 = lean_ctor_get(x_792, 0); +lean_inc(x_808); +x_809 = lean_ctor_get(x_792, 1); +lean_inc(x_809); +lean_dec(x_792); +lean_inc(x_1); +x_810 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_794, x_791, x_1); +x_811 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_811, 0, x_808); +lean_ctor_set(x_811, 1, x_809); +lean_ctor_set(x_811, 2, x_810); +if (lean_is_scalar(x_793)) { + x_812 = lean_alloc_ctor(0, 2, 0); +} else { + x_812 = x_793; +} +lean_ctor_set(x_812, 0, x_1); +lean_ctor_set(x_812, 1, x_811); +x_813 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_813, 0, x_812); +lean_ctor_set(x_813, 1, x_790); +return x_813; +} +} +else +{ +lean_object* x_814; lean_object* x_815; lean_object* x_816; +lean_dec(x_794); +lean_dec(x_791); +lean_dec(x_743); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_660 = !lean_is_exclusive(x_635); -if (x_660 == 0) +x_814 = lean_ctor_get(x_795, 0); +lean_inc(x_814); +lean_dec(x_795); +if (lean_is_scalar(x_793)) { + x_815 = lean_alloc_ctor(0, 2, 0); +} else { + x_815 = x_793; +} +lean_ctor_set(x_815, 0, x_814); +lean_ctor_set(x_815, 1, x_792); +x_816 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_816, 0, x_815); +lean_ctor_set(x_816, 1, x_790); +return x_816; +} +} +} +else +{ +uint8_t x_817; +lean_dec(x_743); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_817 = !lean_is_exclusive(x_744); +if (x_817 == 0) { -return x_635; +return x_744; } else { -lean_object* x_661; lean_object* x_662; lean_object* x_663; -x_661 = lean_ctor_get(x_635, 0); -x_662 = lean_ctor_get(x_635, 1); -lean_inc(x_662); -lean_inc(x_661); -lean_dec(x_635); -x_663 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_663, 0, x_661); -lean_ctor_set(x_663, 1, x_662); -return x_663; +lean_object* x_818; lean_object* x_819; lean_object* x_820; +x_818 = lean_ctor_get(x_744, 0); +x_819 = lean_ctor_get(x_744, 1); +lean_inc(x_819); +lean_inc(x_818); +lean_dec(x_744); +x_820 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_820, 0, x_818); +lean_ctor_set(x_820, 1, x_819); +return x_820; } } } -block_698: +block_875: { -lean_object* x_666; -lean_dec(x_665); +lean_object* x_823; +lean_dec(x_822); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_666 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_666) == 0) -{ -lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; uint8_t x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; -x_667 = lean_ctor_get(x_666, 0); -lean_inc(x_667); -x_668 = lean_ctor_get(x_666, 1); -lean_inc(x_668); -lean_dec(x_666); -x_669 = lean_ctor_get(x_667, 0); -lean_inc(x_669); -lean_dec(x_667); -x_670 = lean_array_get_size(x_3); -x_671 = lean_unsigned_to_nat(0u); -x_672 = lean_unsigned_to_nat(1u); -x_673 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_673, 0, x_671); -lean_ctor_set(x_673, 1, x_670); -lean_ctor_set(x_673, 2, x_672); -x_674 = 0; -x_675 = lean_box(x_674); -x_676 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_676, 0, x_3); -lean_ctor_set(x_676, 1, x_675); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); +x_823 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_823) == 0) +{ +lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; uint8_t x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; +x_824 = lean_ctor_get(x_823, 0); +lean_inc(x_824); +x_825 = lean_ctor_get(x_823, 1); +lean_inc(x_825); +lean_dec(x_823); +x_826 = lean_ctor_get(x_824, 0); +lean_inc(x_826); +lean_dec(x_824); +x_827 = lean_array_get_size(x_3); +x_828 = lean_unsigned_to_nat(0u); +x_829 = lean_unsigned_to_nat(1u); +lean_inc(x_827); +x_830 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_830, 0, x_828); +lean_ctor_set(x_830, 1, x_827); +lean_ctor_set(x_830, 2, x_829); +x_831 = 0; +x_832 = lean_box(x_831); +lean_inc(x_3); +x_833 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_833, 0, x_3); +lean_ctor_set(x_833, 1, x_832); lean_inc(x_2); -x_677 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_669, x_673, x_673, x_676, x_671, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_668); -lean_dec(x_673); -lean_dec(x_669); -if (lean_obj_tag(x_677) == 0) -{ -lean_object* x_678; lean_object* x_679; lean_object* x_680; uint8_t x_681; -x_678 = lean_ctor_get(x_677, 0); -lean_inc(x_678); -x_679 = lean_ctor_get(x_678, 0); -lean_inc(x_679); -x_680 = lean_ctor_get(x_679, 1); -lean_inc(x_680); -x_681 = lean_unbox(x_680); -lean_dec(x_680); -if (x_681 == 0) -{ -lean_object* x_682; lean_object* x_683; lean_object* x_684; -lean_dec(x_679); +x_834 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_826, x_827, x_830, x_830, x_833, x_828, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_825); +lean_dec(x_830); +lean_dec(x_827); +lean_dec(x_826); +lean_dec(x_3); +if (lean_obj_tag(x_834) == 0) +{ +lean_object* x_835; lean_object* x_836; lean_object* x_837; uint8_t x_838; +x_835 = lean_ctor_get(x_834, 0); +lean_inc(x_835); +x_836 = lean_ctor_get(x_835, 0); +lean_inc(x_836); +x_837 = lean_ctor_get(x_836, 1); +lean_inc(x_837); +x_838 = lean_unbox(x_837); +lean_dec(x_837); +if (x_838 == 0) +{ +uint8_t x_839; +lean_dec(x_836); lean_dec(x_2); -x_682 = lean_ctor_get(x_677, 1); -lean_inc(x_682); -lean_dec(x_677); -x_683 = lean_ctor_get(x_678, 1); -lean_inc(x_683); -lean_dec(x_678); -lean_inc(x_1); -x_684 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_683, x_7, x_8, x_9, x_10, x_682); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_684; +x_839 = !lean_is_exclusive(x_834); +if (x_839 == 0) +{ +lean_object* x_840; uint8_t x_841; +x_840 = lean_ctor_get(x_834, 0); +lean_dec(x_840); +x_841 = !lean_is_exclusive(x_835); +if (x_841 == 0) +{ +lean_object* x_842; +x_842 = lean_ctor_get(x_835, 0); +lean_dec(x_842); +lean_ctor_set(x_835, 0, x_1); +return x_834; } else { -lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; -x_685 = lean_ctor_get(x_677, 1); -lean_inc(x_685); -lean_dec(x_677); -x_686 = lean_ctor_get(x_678, 1); -lean_inc(x_686); -lean_dec(x_678); -x_687 = lean_ctor_get(x_679, 0); -lean_inc(x_687); -lean_dec(x_679); -x_688 = l_Lean_mkAppN(x_2, x_687); -lean_dec(x_687); -x_689 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_688, x_5, x_686, x_7, x_8, x_9, x_10, x_685); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_689; +lean_object* x_843; lean_object* x_844; +x_843 = lean_ctor_get(x_835, 1); +lean_inc(x_843); +lean_dec(x_835); +x_844 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_844, 0, x_1); +lean_ctor_set(x_844, 1, x_843); +lean_ctor_set(x_834, 0, x_844); +return x_834; } } else { -uint8_t x_690; +lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; +x_845 = lean_ctor_get(x_834, 1); +lean_inc(x_845); +lean_dec(x_834); +x_846 = lean_ctor_get(x_835, 1); +lean_inc(x_846); +if (lean_is_exclusive(x_835)) { + lean_ctor_release(x_835, 0); + lean_ctor_release(x_835, 1); + x_847 = x_835; +} else { + lean_dec_ref(x_835); + x_847 = lean_box(0); +} +if (lean_is_scalar(x_847)) { + x_848 = lean_alloc_ctor(0, 2, 0); +} else { + x_848 = x_847; +} +lean_ctor_set(x_848, 0, x_1); +lean_ctor_set(x_848, 1, x_846); +x_849 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_849, 0, x_848); +lean_ctor_set(x_849, 1, x_845); +return x_849; +} +} +else +{ +uint8_t x_850; +lean_dec(x_1); +x_850 = !lean_is_exclusive(x_834); +if (x_850 == 0) +{ +lean_object* x_851; uint8_t x_852; +x_851 = lean_ctor_get(x_834, 0); +lean_dec(x_851); +x_852 = !lean_is_exclusive(x_835); +if (x_852 == 0) +{ +lean_object* x_853; lean_object* x_854; lean_object* x_855; +x_853 = lean_ctor_get(x_835, 0); +lean_dec(x_853); +x_854 = lean_ctor_get(x_836, 0); +lean_inc(x_854); +lean_dec(x_836); +x_855 = l_Lean_mkAppN(x_2, x_854); +lean_dec(x_854); +lean_ctor_set(x_835, 0, x_855); +return x_834; +} +else +{ +lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; +x_856 = lean_ctor_get(x_835, 1); +lean_inc(x_856); +lean_dec(x_835); +x_857 = lean_ctor_get(x_836, 0); +lean_inc(x_857); +lean_dec(x_836); +x_858 = l_Lean_mkAppN(x_2, x_857); +lean_dec(x_857); +x_859 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_859, 0, x_858); +lean_ctor_set(x_859, 1, x_856); +lean_ctor_set(x_834, 0, x_859); +return x_834; +} +} +else +{ +lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; +x_860 = lean_ctor_get(x_834, 1); +lean_inc(x_860); +lean_dec(x_834); +x_861 = lean_ctor_get(x_835, 1); +lean_inc(x_861); +if (lean_is_exclusive(x_835)) { + lean_ctor_release(x_835, 0); + lean_ctor_release(x_835, 1); + x_862 = x_835; +} else { + lean_dec_ref(x_835); + x_862 = lean_box(0); +} +x_863 = lean_ctor_get(x_836, 0); +lean_inc(x_863); +lean_dec(x_836); +x_864 = l_Lean_mkAppN(x_2, x_863); +lean_dec(x_863); +if (lean_is_scalar(x_862)) { + x_865 = lean_alloc_ctor(0, 2, 0); +} else { + x_865 = x_862; +} +lean_ctor_set(x_865, 0, x_864); +lean_ctor_set(x_865, 1, x_861); +x_866 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_866, 0, x_865); +lean_ctor_set(x_866, 1, x_860); +return x_866; +} +} +} +else +{ +uint8_t x_867; +lean_dec(x_2); +lean_dec(x_1); +x_867 = !lean_is_exclusive(x_834); +if (x_867 == 0) +{ +return x_834; +} +else +{ +lean_object* x_868; lean_object* x_869; lean_object* x_870; +x_868 = lean_ctor_get(x_834, 0); +x_869 = lean_ctor_get(x_834, 1); +lean_inc(x_869); +lean_inc(x_868); +lean_dec(x_834); +x_870 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_870, 0, x_868); +lean_ctor_set(x_870, 1, x_869); +return x_870; +} +} +} +else +{ +uint8_t x_871; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_690 = !lean_is_exclusive(x_677); -if (x_690 == 0) +x_871 = !lean_is_exclusive(x_823); +if (x_871 == 0) +{ +return x_823; +} +else +{ +lean_object* x_872; lean_object* x_873; lean_object* x_874; +x_872 = lean_ctor_get(x_823, 0); +x_873 = lean_ctor_get(x_823, 1); +lean_inc(x_873); +lean_inc(x_872); +lean_dec(x_823); +x_874 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_874, 0, x_872); +lean_ctor_set(x_874, 1, x_873); +return x_874; +} +} +} +} +case 7: +{ +lean_object* x_888; lean_object* x_967; lean_object* x_1021; uint8_t x_1022; +lean_dec(x_4); +x_1021 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_1022 = l_Lean_Expr_isConstOf(x_2, x_1021); +if (x_1022 == 0) +{ +lean_object* x_1023; +x_1023 = lean_box(0); +x_967 = x_1023; +goto block_1020; +} +else +{ +lean_object* x_1024; lean_object* x_1025; uint8_t x_1026; +x_1024 = lean_array_get_size(x_3); +x_1025 = lean_unsigned_to_nat(2u); +x_1026 = lean_nat_dec_eq(x_1024, x_1025); +if (x_1026 == 0) +{ +lean_object* x_1027; +lean_dec(x_1024); +x_1027 = lean_box(0); +x_967 = x_1027; +goto block_1020; +} +else +{ +lean_object* x_1028; uint8_t x_1029; +x_1028 = lean_unsigned_to_nat(0u); +x_1029 = lean_nat_dec_lt(x_1028, x_1024); +lean_dec(x_1024); +if (x_1029 == 0) +{ +lean_object* x_1030; lean_object* x_1031; +x_1030 = l_Lean_instInhabitedExpr; +x_1031 = l_outOfBounds___rarg(x_1030); +x_888 = x_1031; +goto block_966; +} +else +{ +lean_object* x_1032; +x_1032 = lean_array_fget(x_3, x_1028); +x_888 = x_1032; +goto block_966; +} +} +} +block_966: +{ +lean_object* x_889; +lean_inc(x_888); +x_889 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_888, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_889) == 0) +{ +uint8_t x_890; +x_890 = !lean_is_exclusive(x_889); +if (x_890 == 0) +{ +lean_object* x_891; uint8_t x_892; +x_891 = lean_ctor_get(x_889, 0); +x_892 = !lean_is_exclusive(x_891); +if (x_892 == 0) +{ +lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; +x_893 = lean_ctor_get(x_891, 0); +x_894 = lean_ctor_get(x_891, 1); +x_895 = lean_ctor_get(x_894, 2); +lean_inc(x_895); +lean_inc(x_895); +x_896 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_895, x_893); +if (lean_obj_tag(x_896) == 0) +{ +size_t x_897; size_t x_898; uint8_t x_899; +x_897 = lean_ptr_addr(x_888); +lean_dec(x_888); +x_898 = lean_ptr_addr(x_893); +x_899 = lean_usize_dec_eq(x_897, x_898); +if (x_899 == 0) +{ +lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; +lean_dec(x_1); +x_900 = lean_unsigned_to_nat(0u); +lean_inc(x_893); +x_901 = lean_array_set(x_3, x_900, x_893); +x_902 = l_Lean_mkAppN(x_2, x_901); +lean_dec(x_901); +x_903 = lean_ctor_get(x_894, 0); +lean_inc(x_903); +x_904 = lean_ctor_get(x_894, 1); +lean_inc(x_904); +lean_dec(x_894); +lean_inc(x_902); +x_905 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_895, x_893, x_902); +x_906 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_906, 0, x_903); +lean_ctor_set(x_906, 1, x_904); +lean_ctor_set(x_906, 2, x_905); +lean_ctor_set(x_891, 1, x_906); +lean_ctor_set(x_891, 0, x_902); +return x_889; +} +else +{ +lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; +lean_dec(x_3); +lean_dec(x_2); +x_907 = lean_ctor_get(x_894, 0); +lean_inc(x_907); +x_908 = lean_ctor_get(x_894, 1); +lean_inc(x_908); +lean_dec(x_894); +lean_inc(x_1); +x_909 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_895, x_893, x_1); +x_910 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_910, 0, x_907); +lean_ctor_set(x_910, 1, x_908); +lean_ctor_set(x_910, 2, x_909); +lean_ctor_set(x_891, 1, x_910); +lean_ctor_set(x_891, 0, x_1); +return x_889; +} +} +else +{ +lean_object* x_911; +lean_dec(x_895); +lean_dec(x_893); +lean_dec(x_888); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_911 = lean_ctor_get(x_896, 0); +lean_inc(x_911); +lean_dec(x_896); +lean_ctor_set(x_891, 0, x_911); +return x_889; +} +} +else +{ +lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; +x_912 = lean_ctor_get(x_891, 0); +x_913 = lean_ctor_get(x_891, 1); +lean_inc(x_913); +lean_inc(x_912); +lean_dec(x_891); +x_914 = lean_ctor_get(x_913, 2); +lean_inc(x_914); +lean_inc(x_914); +x_915 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_914, x_912); +if (lean_obj_tag(x_915) == 0) +{ +size_t x_916; size_t x_917; uint8_t x_918; +x_916 = lean_ptr_addr(x_888); +lean_dec(x_888); +x_917 = lean_ptr_addr(x_912); +x_918 = lean_usize_dec_eq(x_916, x_917); +if (x_918 == 0) +{ +lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; +lean_dec(x_1); +x_919 = lean_unsigned_to_nat(0u); +lean_inc(x_912); +x_920 = lean_array_set(x_3, x_919, x_912); +x_921 = l_Lean_mkAppN(x_2, x_920); +lean_dec(x_920); +x_922 = lean_ctor_get(x_913, 0); +lean_inc(x_922); +x_923 = lean_ctor_get(x_913, 1); +lean_inc(x_923); +lean_dec(x_913); +lean_inc(x_921); +x_924 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_914, x_912, x_921); +x_925 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_925, 0, x_922); +lean_ctor_set(x_925, 1, x_923); +lean_ctor_set(x_925, 2, x_924); +x_926 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_926, 0, x_921); +lean_ctor_set(x_926, 1, x_925); +lean_ctor_set(x_889, 0, x_926); +return x_889; +} +else +{ +lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; +lean_dec(x_3); +lean_dec(x_2); +x_927 = lean_ctor_get(x_913, 0); +lean_inc(x_927); +x_928 = lean_ctor_get(x_913, 1); +lean_inc(x_928); +lean_dec(x_913); +lean_inc(x_1); +x_929 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_914, x_912, x_1); +x_930 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_930, 0, x_927); +lean_ctor_set(x_930, 1, x_928); +lean_ctor_set(x_930, 2, x_929); +x_931 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_931, 0, x_1); +lean_ctor_set(x_931, 1, x_930); +lean_ctor_set(x_889, 0, x_931); +return x_889; +} +} +else +{ +lean_object* x_932; lean_object* x_933; +lean_dec(x_914); +lean_dec(x_912); +lean_dec(x_888); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_932 = lean_ctor_get(x_915, 0); +lean_inc(x_932); +lean_dec(x_915); +x_933 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_933, 0, x_932); +lean_ctor_set(x_933, 1, x_913); +lean_ctor_set(x_889, 0, x_933); +return x_889; +} +} +} +else +{ +lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; +x_934 = lean_ctor_get(x_889, 0); +x_935 = lean_ctor_get(x_889, 1); +lean_inc(x_935); +lean_inc(x_934); +lean_dec(x_889); +x_936 = lean_ctor_get(x_934, 0); +lean_inc(x_936); +x_937 = lean_ctor_get(x_934, 1); +lean_inc(x_937); +if (lean_is_exclusive(x_934)) { + lean_ctor_release(x_934, 0); + lean_ctor_release(x_934, 1); + x_938 = x_934; +} else { + lean_dec_ref(x_934); + x_938 = lean_box(0); +} +x_939 = lean_ctor_get(x_937, 2); +lean_inc(x_939); +lean_inc(x_939); +x_940 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_939, x_936); +if (lean_obj_tag(x_940) == 0) +{ +size_t x_941; size_t x_942; uint8_t x_943; +x_941 = lean_ptr_addr(x_888); +lean_dec(x_888); +x_942 = lean_ptr_addr(x_936); +x_943 = lean_usize_dec_eq(x_941, x_942); +if (x_943 == 0) +{ +lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; +lean_dec(x_1); +x_944 = lean_unsigned_to_nat(0u); +lean_inc(x_936); +x_945 = lean_array_set(x_3, x_944, x_936); +x_946 = l_Lean_mkAppN(x_2, x_945); +lean_dec(x_945); +x_947 = lean_ctor_get(x_937, 0); +lean_inc(x_947); +x_948 = lean_ctor_get(x_937, 1); +lean_inc(x_948); +lean_dec(x_937); +lean_inc(x_946); +x_949 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_939, x_936, x_946); +x_950 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_950, 0, x_947); +lean_ctor_set(x_950, 1, x_948); +lean_ctor_set(x_950, 2, x_949); +if (lean_is_scalar(x_938)) { + x_951 = lean_alloc_ctor(0, 2, 0); +} else { + x_951 = x_938; +} +lean_ctor_set(x_951, 0, x_946); +lean_ctor_set(x_951, 1, x_950); +x_952 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_952, 0, x_951); +lean_ctor_set(x_952, 1, x_935); +return x_952; +} +else +{ +lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; +lean_dec(x_3); +lean_dec(x_2); +x_953 = lean_ctor_get(x_937, 0); +lean_inc(x_953); +x_954 = lean_ctor_get(x_937, 1); +lean_inc(x_954); +lean_dec(x_937); +lean_inc(x_1); +x_955 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_939, x_936, x_1); +x_956 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_956, 0, x_953); +lean_ctor_set(x_956, 1, x_954); +lean_ctor_set(x_956, 2, x_955); +if (lean_is_scalar(x_938)) { + x_957 = lean_alloc_ctor(0, 2, 0); +} else { + x_957 = x_938; +} +lean_ctor_set(x_957, 0, x_1); +lean_ctor_set(x_957, 1, x_956); +x_958 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_958, 0, x_957); +lean_ctor_set(x_958, 1, x_935); +return x_958; +} +} +else +{ +lean_object* x_959; lean_object* x_960; lean_object* x_961; +lean_dec(x_939); +lean_dec(x_936); +lean_dec(x_888); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_959 = lean_ctor_get(x_940, 0); +lean_inc(x_959); +lean_dec(x_940); +if (lean_is_scalar(x_938)) { + x_960 = lean_alloc_ctor(0, 2, 0); +} else { + x_960 = x_938; +} +lean_ctor_set(x_960, 0, x_959); +lean_ctor_set(x_960, 1, x_937); +x_961 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_961, 0, x_960); +lean_ctor_set(x_961, 1, x_935); +return x_961; +} +} +} +else +{ +uint8_t x_962; +lean_dec(x_888); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_962 = !lean_is_exclusive(x_889); +if (x_962 == 0) +{ +return x_889; +} +else +{ +lean_object* x_963; lean_object* x_964; lean_object* x_965; +x_963 = lean_ctor_get(x_889, 0); +x_964 = lean_ctor_get(x_889, 1); +lean_inc(x_964); +lean_inc(x_963); +lean_dec(x_889); +x_965 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_965, 0, x_963); +lean_ctor_set(x_965, 1, x_964); +return x_965; +} +} +} +block_1020: +{ +lean_object* x_968; +lean_dec(x_967); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_968 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_968) == 0) +{ +lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; uint8_t x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; +x_969 = lean_ctor_get(x_968, 0); +lean_inc(x_969); +x_970 = lean_ctor_get(x_968, 1); +lean_inc(x_970); +lean_dec(x_968); +x_971 = lean_ctor_get(x_969, 0); +lean_inc(x_971); +lean_dec(x_969); +x_972 = lean_array_get_size(x_3); +x_973 = lean_unsigned_to_nat(0u); +x_974 = lean_unsigned_to_nat(1u); +lean_inc(x_972); +x_975 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_975, 0, x_973); +lean_ctor_set(x_975, 1, x_972); +lean_ctor_set(x_975, 2, x_974); +x_976 = 0; +x_977 = lean_box(x_976); +lean_inc(x_3); +x_978 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_978, 0, x_3); +lean_ctor_set(x_978, 1, x_977); +lean_inc(x_2); +x_979 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_971, x_972, x_975, x_975, x_978, x_973, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_970); +lean_dec(x_975); +lean_dec(x_972); +lean_dec(x_971); +lean_dec(x_3); +if (lean_obj_tag(x_979) == 0) +{ +lean_object* x_980; lean_object* x_981; lean_object* x_982; uint8_t x_983; +x_980 = lean_ctor_get(x_979, 0); +lean_inc(x_980); +x_981 = lean_ctor_get(x_980, 0); +lean_inc(x_981); +x_982 = lean_ctor_get(x_981, 1); +lean_inc(x_982); +x_983 = lean_unbox(x_982); +lean_dec(x_982); +if (x_983 == 0) +{ +uint8_t x_984; +lean_dec(x_981); +lean_dec(x_2); +x_984 = !lean_is_exclusive(x_979); +if (x_984 == 0) +{ +lean_object* x_985; uint8_t x_986; +x_985 = lean_ctor_get(x_979, 0); +lean_dec(x_985); +x_986 = !lean_is_exclusive(x_980); +if (x_986 == 0) +{ +lean_object* x_987; +x_987 = lean_ctor_get(x_980, 0); +lean_dec(x_987); +lean_ctor_set(x_980, 0, x_1); +return x_979; +} +else +{ +lean_object* x_988; lean_object* x_989; +x_988 = lean_ctor_get(x_980, 1); +lean_inc(x_988); +lean_dec(x_980); +x_989 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_989, 0, x_1); +lean_ctor_set(x_989, 1, x_988); +lean_ctor_set(x_979, 0, x_989); +return x_979; +} +} +else +{ +lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; +x_990 = lean_ctor_get(x_979, 1); +lean_inc(x_990); +lean_dec(x_979); +x_991 = lean_ctor_get(x_980, 1); +lean_inc(x_991); +if (lean_is_exclusive(x_980)) { + lean_ctor_release(x_980, 0); + lean_ctor_release(x_980, 1); + x_992 = x_980; +} else { + lean_dec_ref(x_980); + x_992 = lean_box(0); +} +if (lean_is_scalar(x_992)) { + x_993 = lean_alloc_ctor(0, 2, 0); +} else { + x_993 = x_992; +} +lean_ctor_set(x_993, 0, x_1); +lean_ctor_set(x_993, 1, x_991); +x_994 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_994, 0, x_993); +lean_ctor_set(x_994, 1, x_990); +return x_994; +} +} +else +{ +uint8_t x_995; +lean_dec(x_1); +x_995 = !lean_is_exclusive(x_979); +if (x_995 == 0) +{ +lean_object* x_996; uint8_t x_997; +x_996 = lean_ctor_get(x_979, 0); +lean_dec(x_996); +x_997 = !lean_is_exclusive(x_980); +if (x_997 == 0) +{ +lean_object* x_998; lean_object* x_999; lean_object* x_1000; +x_998 = lean_ctor_get(x_980, 0); +lean_dec(x_998); +x_999 = lean_ctor_get(x_981, 0); +lean_inc(x_999); +lean_dec(x_981); +x_1000 = l_Lean_mkAppN(x_2, x_999); +lean_dec(x_999); +lean_ctor_set(x_980, 0, x_1000); +return x_979; +} +else +{ +lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; +x_1001 = lean_ctor_get(x_980, 1); +lean_inc(x_1001); +lean_dec(x_980); +x_1002 = lean_ctor_get(x_981, 0); +lean_inc(x_1002); +lean_dec(x_981); +x_1003 = l_Lean_mkAppN(x_2, x_1002); +lean_dec(x_1002); +x_1004 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1004, 0, x_1003); +lean_ctor_set(x_1004, 1, x_1001); +lean_ctor_set(x_979, 0, x_1004); +return x_979; +} +} +else +{ +lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; +x_1005 = lean_ctor_get(x_979, 1); +lean_inc(x_1005); +lean_dec(x_979); +x_1006 = lean_ctor_get(x_980, 1); +lean_inc(x_1006); +if (lean_is_exclusive(x_980)) { + lean_ctor_release(x_980, 0); + lean_ctor_release(x_980, 1); + x_1007 = x_980; +} else { + lean_dec_ref(x_980); + x_1007 = lean_box(0); +} +x_1008 = lean_ctor_get(x_981, 0); +lean_inc(x_1008); +lean_dec(x_981); +x_1009 = l_Lean_mkAppN(x_2, x_1008); +lean_dec(x_1008); +if (lean_is_scalar(x_1007)) { + x_1010 = lean_alloc_ctor(0, 2, 0); +} else { + x_1010 = x_1007; +} +lean_ctor_set(x_1010, 0, x_1009); +lean_ctor_set(x_1010, 1, x_1006); +x_1011 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1011, 0, x_1010); +lean_ctor_set(x_1011, 1, x_1005); +return x_1011; +} +} +} +else +{ +uint8_t x_1012; +lean_dec(x_2); +lean_dec(x_1); +x_1012 = !lean_is_exclusive(x_979); +if (x_1012 == 0) +{ +return x_979; +} +else +{ +lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; +x_1013 = lean_ctor_get(x_979, 0); +x_1014 = lean_ctor_get(x_979, 1); +lean_inc(x_1014); +lean_inc(x_1013); +lean_dec(x_979); +x_1015 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1015, 0, x_1013); +lean_ctor_set(x_1015, 1, x_1014); +return x_1015; +} +} +} +else +{ +uint8_t x_1016; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1016 = !lean_is_exclusive(x_968); +if (x_1016 == 0) +{ +return x_968; +} +else +{ +lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; +x_1017 = lean_ctor_get(x_968, 0); +x_1018 = lean_ctor_get(x_968, 1); +lean_inc(x_1018); +lean_inc(x_1017); +lean_dec(x_968); +x_1019 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1019, 0, x_1017); +lean_ctor_set(x_1019, 1, x_1018); +return x_1019; +} +} +} +} +case 8: +{ +lean_object* x_1033; lean_object* x_1112; lean_object* x_1166; uint8_t x_1167; +lean_dec(x_4); +x_1166 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_1167 = l_Lean_Expr_isConstOf(x_2, x_1166); +if (x_1167 == 0) +{ +lean_object* x_1168; +x_1168 = lean_box(0); +x_1112 = x_1168; +goto block_1165; +} +else +{ +lean_object* x_1169; lean_object* x_1170; uint8_t x_1171; +x_1169 = lean_array_get_size(x_3); +x_1170 = lean_unsigned_to_nat(2u); +x_1171 = lean_nat_dec_eq(x_1169, x_1170); +if (x_1171 == 0) +{ +lean_object* x_1172; +lean_dec(x_1169); +x_1172 = lean_box(0); +x_1112 = x_1172; +goto block_1165; +} +else +{ +lean_object* x_1173; uint8_t x_1174; +x_1173 = lean_unsigned_to_nat(0u); +x_1174 = lean_nat_dec_lt(x_1173, x_1169); +lean_dec(x_1169); +if (x_1174 == 0) +{ +lean_object* x_1175; lean_object* x_1176; +x_1175 = l_Lean_instInhabitedExpr; +x_1176 = l_outOfBounds___rarg(x_1175); +x_1033 = x_1176; +goto block_1111; +} +else +{ +lean_object* x_1177; +x_1177 = lean_array_fget(x_3, x_1173); +x_1033 = x_1177; +goto block_1111; +} +} +} +block_1111: +{ +lean_object* x_1034; +lean_inc(x_1033); +x_1034 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_1033, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_1034) == 0) +{ +uint8_t x_1035; +x_1035 = !lean_is_exclusive(x_1034); +if (x_1035 == 0) +{ +lean_object* x_1036; uint8_t x_1037; +x_1036 = lean_ctor_get(x_1034, 0); +x_1037 = !lean_is_exclusive(x_1036); +if (x_1037 == 0) +{ +lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; +x_1038 = lean_ctor_get(x_1036, 0); +x_1039 = lean_ctor_get(x_1036, 1); +x_1040 = lean_ctor_get(x_1039, 2); +lean_inc(x_1040); +lean_inc(x_1040); +x_1041 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1040, x_1038); +if (lean_obj_tag(x_1041) == 0) +{ +size_t x_1042; size_t x_1043; uint8_t x_1044; +x_1042 = lean_ptr_addr(x_1033); +lean_dec(x_1033); +x_1043 = lean_ptr_addr(x_1038); +x_1044 = lean_usize_dec_eq(x_1042, x_1043); +if (x_1044 == 0) +{ +lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; +lean_dec(x_1); +x_1045 = lean_unsigned_to_nat(0u); +lean_inc(x_1038); +x_1046 = lean_array_set(x_3, x_1045, x_1038); +x_1047 = l_Lean_mkAppN(x_2, x_1046); +lean_dec(x_1046); +x_1048 = lean_ctor_get(x_1039, 0); +lean_inc(x_1048); +x_1049 = lean_ctor_get(x_1039, 1); +lean_inc(x_1049); +lean_dec(x_1039); +lean_inc(x_1047); +x_1050 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1040, x_1038, x_1047); +x_1051 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1051, 0, x_1048); +lean_ctor_set(x_1051, 1, x_1049); +lean_ctor_set(x_1051, 2, x_1050); +lean_ctor_set(x_1036, 1, x_1051); +lean_ctor_set(x_1036, 0, x_1047); +return x_1034; +} +else +{ +lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; +lean_dec(x_3); +lean_dec(x_2); +x_1052 = lean_ctor_get(x_1039, 0); +lean_inc(x_1052); +x_1053 = lean_ctor_get(x_1039, 1); +lean_inc(x_1053); +lean_dec(x_1039); +lean_inc(x_1); +x_1054 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1040, x_1038, x_1); +x_1055 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1055, 0, x_1052); +lean_ctor_set(x_1055, 1, x_1053); +lean_ctor_set(x_1055, 2, x_1054); +lean_ctor_set(x_1036, 1, x_1055); +lean_ctor_set(x_1036, 0, x_1); +return x_1034; +} +} +else +{ +lean_object* x_1056; +lean_dec(x_1040); +lean_dec(x_1038); +lean_dec(x_1033); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1056 = lean_ctor_get(x_1041, 0); +lean_inc(x_1056); +lean_dec(x_1041); +lean_ctor_set(x_1036, 0, x_1056); +return x_1034; +} +} +else +{ +lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; +x_1057 = lean_ctor_get(x_1036, 0); +x_1058 = lean_ctor_get(x_1036, 1); +lean_inc(x_1058); +lean_inc(x_1057); +lean_dec(x_1036); +x_1059 = lean_ctor_get(x_1058, 2); +lean_inc(x_1059); +lean_inc(x_1059); +x_1060 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1059, x_1057); +if (lean_obj_tag(x_1060) == 0) +{ +size_t x_1061; size_t x_1062; uint8_t x_1063; +x_1061 = lean_ptr_addr(x_1033); +lean_dec(x_1033); +x_1062 = lean_ptr_addr(x_1057); +x_1063 = lean_usize_dec_eq(x_1061, x_1062); +if (x_1063 == 0) +{ +lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; +lean_dec(x_1); +x_1064 = lean_unsigned_to_nat(0u); +lean_inc(x_1057); +x_1065 = lean_array_set(x_3, x_1064, x_1057); +x_1066 = l_Lean_mkAppN(x_2, x_1065); +lean_dec(x_1065); +x_1067 = lean_ctor_get(x_1058, 0); +lean_inc(x_1067); +x_1068 = lean_ctor_get(x_1058, 1); +lean_inc(x_1068); +lean_dec(x_1058); +lean_inc(x_1066); +x_1069 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1059, x_1057, x_1066); +x_1070 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1070, 0, x_1067); +lean_ctor_set(x_1070, 1, x_1068); +lean_ctor_set(x_1070, 2, x_1069); +x_1071 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1071, 0, x_1066); +lean_ctor_set(x_1071, 1, x_1070); +lean_ctor_set(x_1034, 0, x_1071); +return x_1034; +} +else +{ +lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; +lean_dec(x_3); +lean_dec(x_2); +x_1072 = lean_ctor_get(x_1058, 0); +lean_inc(x_1072); +x_1073 = lean_ctor_get(x_1058, 1); +lean_inc(x_1073); +lean_dec(x_1058); +lean_inc(x_1); +x_1074 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1059, x_1057, x_1); +x_1075 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1075, 0, x_1072); +lean_ctor_set(x_1075, 1, x_1073); +lean_ctor_set(x_1075, 2, x_1074); +x_1076 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1076, 0, x_1); +lean_ctor_set(x_1076, 1, x_1075); +lean_ctor_set(x_1034, 0, x_1076); +return x_1034; +} +} +else +{ +lean_object* x_1077; lean_object* x_1078; +lean_dec(x_1059); +lean_dec(x_1057); +lean_dec(x_1033); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1077 = lean_ctor_get(x_1060, 0); +lean_inc(x_1077); +lean_dec(x_1060); +x_1078 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1078, 0, x_1077); +lean_ctor_set(x_1078, 1, x_1058); +lean_ctor_set(x_1034, 0, x_1078); +return x_1034; +} +} +} +else +{ +lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; +x_1079 = lean_ctor_get(x_1034, 0); +x_1080 = lean_ctor_get(x_1034, 1); +lean_inc(x_1080); +lean_inc(x_1079); +lean_dec(x_1034); +x_1081 = lean_ctor_get(x_1079, 0); +lean_inc(x_1081); +x_1082 = lean_ctor_get(x_1079, 1); +lean_inc(x_1082); +if (lean_is_exclusive(x_1079)) { + lean_ctor_release(x_1079, 0); + lean_ctor_release(x_1079, 1); + x_1083 = x_1079; +} else { + lean_dec_ref(x_1079); + x_1083 = lean_box(0); +} +x_1084 = lean_ctor_get(x_1082, 2); +lean_inc(x_1084); +lean_inc(x_1084); +x_1085 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1084, x_1081); +if (lean_obj_tag(x_1085) == 0) +{ +size_t x_1086; size_t x_1087; uint8_t x_1088; +x_1086 = lean_ptr_addr(x_1033); +lean_dec(x_1033); +x_1087 = lean_ptr_addr(x_1081); +x_1088 = lean_usize_dec_eq(x_1086, x_1087); +if (x_1088 == 0) +{ +lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; +lean_dec(x_1); +x_1089 = lean_unsigned_to_nat(0u); +lean_inc(x_1081); +x_1090 = lean_array_set(x_3, x_1089, x_1081); +x_1091 = l_Lean_mkAppN(x_2, x_1090); +lean_dec(x_1090); +x_1092 = lean_ctor_get(x_1082, 0); +lean_inc(x_1092); +x_1093 = lean_ctor_get(x_1082, 1); +lean_inc(x_1093); +lean_dec(x_1082); +lean_inc(x_1091); +x_1094 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1084, x_1081, x_1091); +x_1095 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1095, 0, x_1092); +lean_ctor_set(x_1095, 1, x_1093); +lean_ctor_set(x_1095, 2, x_1094); +if (lean_is_scalar(x_1083)) { + x_1096 = lean_alloc_ctor(0, 2, 0); +} else { + x_1096 = x_1083; +} +lean_ctor_set(x_1096, 0, x_1091); +lean_ctor_set(x_1096, 1, x_1095); +x_1097 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1097, 0, x_1096); +lean_ctor_set(x_1097, 1, x_1080); +return x_1097; +} +else +{ +lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; +lean_dec(x_3); +lean_dec(x_2); +x_1098 = lean_ctor_get(x_1082, 0); +lean_inc(x_1098); +x_1099 = lean_ctor_get(x_1082, 1); +lean_inc(x_1099); +lean_dec(x_1082); +lean_inc(x_1); +x_1100 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1084, x_1081, x_1); +x_1101 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1101, 0, x_1098); +lean_ctor_set(x_1101, 1, x_1099); +lean_ctor_set(x_1101, 2, x_1100); +if (lean_is_scalar(x_1083)) { + x_1102 = lean_alloc_ctor(0, 2, 0); +} else { + x_1102 = x_1083; +} +lean_ctor_set(x_1102, 0, x_1); +lean_ctor_set(x_1102, 1, x_1101); +x_1103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1103, 0, x_1102); +lean_ctor_set(x_1103, 1, x_1080); +return x_1103; +} +} +else +{ +lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; +lean_dec(x_1084); +lean_dec(x_1081); +lean_dec(x_1033); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1104 = lean_ctor_get(x_1085, 0); +lean_inc(x_1104); +lean_dec(x_1085); +if (lean_is_scalar(x_1083)) { + x_1105 = lean_alloc_ctor(0, 2, 0); +} else { + x_1105 = x_1083; +} +lean_ctor_set(x_1105, 0, x_1104); +lean_ctor_set(x_1105, 1, x_1082); +x_1106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1106, 0, x_1105); +lean_ctor_set(x_1106, 1, x_1080); +return x_1106; +} +} +} +else +{ +uint8_t x_1107; +lean_dec(x_1033); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1107 = !lean_is_exclusive(x_1034); +if (x_1107 == 0) +{ +return x_1034; +} +else +{ +lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; +x_1108 = lean_ctor_get(x_1034, 0); +x_1109 = lean_ctor_get(x_1034, 1); +lean_inc(x_1109); +lean_inc(x_1108); +lean_dec(x_1034); +x_1110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1110, 0, x_1108); +lean_ctor_set(x_1110, 1, x_1109); +return x_1110; +} +} +} +block_1165: +{ +lean_object* x_1113; +lean_dec(x_1112); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_1113 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_1113) == 0) +{ +lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; uint8_t x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; +x_1114 = lean_ctor_get(x_1113, 0); +lean_inc(x_1114); +x_1115 = lean_ctor_get(x_1113, 1); +lean_inc(x_1115); +lean_dec(x_1113); +x_1116 = lean_ctor_get(x_1114, 0); +lean_inc(x_1116); +lean_dec(x_1114); +x_1117 = lean_array_get_size(x_3); +x_1118 = lean_unsigned_to_nat(0u); +x_1119 = lean_unsigned_to_nat(1u); +lean_inc(x_1117); +x_1120 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1120, 0, x_1118); +lean_ctor_set(x_1120, 1, x_1117); +lean_ctor_set(x_1120, 2, x_1119); +x_1121 = 0; +x_1122 = lean_box(x_1121); +lean_inc(x_3); +x_1123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1123, 0, x_3); +lean_ctor_set(x_1123, 1, x_1122); +lean_inc(x_2); +x_1124 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_1116, x_1117, x_1120, x_1120, x_1123, x_1118, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_1115); +lean_dec(x_1120); +lean_dec(x_1117); +lean_dec(x_1116); +lean_dec(x_3); +if (lean_obj_tag(x_1124) == 0) +{ +lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; uint8_t x_1128; +x_1125 = lean_ctor_get(x_1124, 0); +lean_inc(x_1125); +x_1126 = lean_ctor_get(x_1125, 0); +lean_inc(x_1126); +x_1127 = lean_ctor_get(x_1126, 1); +lean_inc(x_1127); +x_1128 = lean_unbox(x_1127); +lean_dec(x_1127); +if (x_1128 == 0) +{ +uint8_t x_1129; +lean_dec(x_1126); +lean_dec(x_2); +x_1129 = !lean_is_exclusive(x_1124); +if (x_1129 == 0) +{ +lean_object* x_1130; uint8_t x_1131; +x_1130 = lean_ctor_get(x_1124, 0); +lean_dec(x_1130); +x_1131 = !lean_is_exclusive(x_1125); +if (x_1131 == 0) +{ +lean_object* x_1132; +x_1132 = lean_ctor_get(x_1125, 0); +lean_dec(x_1132); +lean_ctor_set(x_1125, 0, x_1); +return x_1124; +} +else +{ +lean_object* x_1133; lean_object* x_1134; +x_1133 = lean_ctor_get(x_1125, 1); +lean_inc(x_1133); +lean_dec(x_1125); +x_1134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1134, 0, x_1); +lean_ctor_set(x_1134, 1, x_1133); +lean_ctor_set(x_1124, 0, x_1134); +return x_1124; +} +} +else +{ +lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; +x_1135 = lean_ctor_get(x_1124, 1); +lean_inc(x_1135); +lean_dec(x_1124); +x_1136 = lean_ctor_get(x_1125, 1); +lean_inc(x_1136); +if (lean_is_exclusive(x_1125)) { + lean_ctor_release(x_1125, 0); + lean_ctor_release(x_1125, 1); + x_1137 = x_1125; +} else { + lean_dec_ref(x_1125); + x_1137 = lean_box(0); +} +if (lean_is_scalar(x_1137)) { + x_1138 = lean_alloc_ctor(0, 2, 0); +} else { + x_1138 = x_1137; +} +lean_ctor_set(x_1138, 0, x_1); +lean_ctor_set(x_1138, 1, x_1136); +x_1139 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1139, 0, x_1138); +lean_ctor_set(x_1139, 1, x_1135); +return x_1139; +} +} +else +{ +uint8_t x_1140; +lean_dec(x_1); +x_1140 = !lean_is_exclusive(x_1124); +if (x_1140 == 0) +{ +lean_object* x_1141; uint8_t x_1142; +x_1141 = lean_ctor_get(x_1124, 0); +lean_dec(x_1141); +x_1142 = !lean_is_exclusive(x_1125); +if (x_1142 == 0) +{ +lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; +x_1143 = lean_ctor_get(x_1125, 0); +lean_dec(x_1143); +x_1144 = lean_ctor_get(x_1126, 0); +lean_inc(x_1144); +lean_dec(x_1126); +x_1145 = l_Lean_mkAppN(x_2, x_1144); +lean_dec(x_1144); +lean_ctor_set(x_1125, 0, x_1145); +return x_1124; +} +else +{ +lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; +x_1146 = lean_ctor_get(x_1125, 1); +lean_inc(x_1146); +lean_dec(x_1125); +x_1147 = lean_ctor_get(x_1126, 0); +lean_inc(x_1147); +lean_dec(x_1126); +x_1148 = l_Lean_mkAppN(x_2, x_1147); +lean_dec(x_1147); +x_1149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1149, 0, x_1148); +lean_ctor_set(x_1149, 1, x_1146); +lean_ctor_set(x_1124, 0, x_1149); +return x_1124; +} +} +else +{ +lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; +x_1150 = lean_ctor_get(x_1124, 1); +lean_inc(x_1150); +lean_dec(x_1124); +x_1151 = lean_ctor_get(x_1125, 1); +lean_inc(x_1151); +if (lean_is_exclusive(x_1125)) { + lean_ctor_release(x_1125, 0); + lean_ctor_release(x_1125, 1); + x_1152 = x_1125; +} else { + lean_dec_ref(x_1125); + x_1152 = lean_box(0); +} +x_1153 = lean_ctor_get(x_1126, 0); +lean_inc(x_1153); +lean_dec(x_1126); +x_1154 = l_Lean_mkAppN(x_2, x_1153); +lean_dec(x_1153); +if (lean_is_scalar(x_1152)) { + x_1155 = lean_alloc_ctor(0, 2, 0); +} else { + x_1155 = x_1152; +} +lean_ctor_set(x_1155, 0, x_1154); +lean_ctor_set(x_1155, 1, x_1151); +x_1156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1156, 0, x_1155); +lean_ctor_set(x_1156, 1, x_1150); +return x_1156; +} +} +} +else +{ +uint8_t x_1157; +lean_dec(x_2); +lean_dec(x_1); +x_1157 = !lean_is_exclusive(x_1124); +if (x_1157 == 0) +{ +return x_1124; +} +else +{ +lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; +x_1158 = lean_ctor_get(x_1124, 0); +x_1159 = lean_ctor_get(x_1124, 1); +lean_inc(x_1159); +lean_inc(x_1158); +lean_dec(x_1124); +x_1160 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1160, 0, x_1158); +lean_ctor_set(x_1160, 1, x_1159); +return x_1160; +} +} +} +else +{ +uint8_t x_1161; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1161 = !lean_is_exclusive(x_1113); +if (x_1161 == 0) +{ +return x_1113; +} +else +{ +lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; +x_1162 = lean_ctor_get(x_1113, 0); +x_1163 = lean_ctor_get(x_1113, 1); +lean_inc(x_1163); +lean_inc(x_1162); +lean_dec(x_1113); +x_1164 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1164, 0, x_1162); +lean_ctor_set(x_1164, 1, x_1163); +return x_1164; +} +} +} +} +case 9: +{ +lean_object* x_1178; lean_object* x_1257; lean_object* x_1311; uint8_t x_1312; +lean_dec(x_4); +x_1311 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_1312 = l_Lean_Expr_isConstOf(x_2, x_1311); +if (x_1312 == 0) +{ +lean_object* x_1313; +x_1313 = lean_box(0); +x_1257 = x_1313; +goto block_1310; +} +else +{ +lean_object* x_1314; lean_object* x_1315; uint8_t x_1316; +x_1314 = lean_array_get_size(x_3); +x_1315 = lean_unsigned_to_nat(2u); +x_1316 = lean_nat_dec_eq(x_1314, x_1315); +if (x_1316 == 0) +{ +lean_object* x_1317; +lean_dec(x_1314); +x_1317 = lean_box(0); +x_1257 = x_1317; +goto block_1310; +} +else +{ +lean_object* x_1318; uint8_t x_1319; +x_1318 = lean_unsigned_to_nat(0u); +x_1319 = lean_nat_dec_lt(x_1318, x_1314); +lean_dec(x_1314); +if (x_1319 == 0) +{ +lean_object* x_1320; lean_object* x_1321; +x_1320 = l_Lean_instInhabitedExpr; +x_1321 = l_outOfBounds___rarg(x_1320); +x_1178 = x_1321; +goto block_1256; +} +else +{ +lean_object* x_1322; +x_1322 = lean_array_fget(x_3, x_1318); +x_1178 = x_1322; +goto block_1256; +} +} +} +block_1256: +{ +lean_object* x_1179; +lean_inc(x_1178); +x_1179 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_1178, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_1179) == 0) +{ +uint8_t x_1180; +x_1180 = !lean_is_exclusive(x_1179); +if (x_1180 == 0) +{ +lean_object* x_1181; uint8_t x_1182; +x_1181 = lean_ctor_get(x_1179, 0); +x_1182 = !lean_is_exclusive(x_1181); +if (x_1182 == 0) +{ +lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; +x_1183 = lean_ctor_get(x_1181, 0); +x_1184 = lean_ctor_get(x_1181, 1); +x_1185 = lean_ctor_get(x_1184, 2); +lean_inc(x_1185); +lean_inc(x_1185); +x_1186 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1185, x_1183); +if (lean_obj_tag(x_1186) == 0) +{ +size_t x_1187; size_t x_1188; uint8_t x_1189; +x_1187 = lean_ptr_addr(x_1178); +lean_dec(x_1178); +x_1188 = lean_ptr_addr(x_1183); +x_1189 = lean_usize_dec_eq(x_1187, x_1188); +if (x_1189 == 0) +{ +lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; +lean_dec(x_1); +x_1190 = lean_unsigned_to_nat(0u); +lean_inc(x_1183); +x_1191 = lean_array_set(x_3, x_1190, x_1183); +x_1192 = l_Lean_mkAppN(x_2, x_1191); +lean_dec(x_1191); +x_1193 = lean_ctor_get(x_1184, 0); +lean_inc(x_1193); +x_1194 = lean_ctor_get(x_1184, 1); +lean_inc(x_1194); +lean_dec(x_1184); +lean_inc(x_1192); +x_1195 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1185, x_1183, x_1192); +x_1196 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1196, 0, x_1193); +lean_ctor_set(x_1196, 1, x_1194); +lean_ctor_set(x_1196, 2, x_1195); +lean_ctor_set(x_1181, 1, x_1196); +lean_ctor_set(x_1181, 0, x_1192); +return x_1179; +} +else +{ +lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; +lean_dec(x_3); +lean_dec(x_2); +x_1197 = lean_ctor_get(x_1184, 0); +lean_inc(x_1197); +x_1198 = lean_ctor_get(x_1184, 1); +lean_inc(x_1198); +lean_dec(x_1184); +lean_inc(x_1); +x_1199 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1185, x_1183, x_1); +x_1200 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1200, 0, x_1197); +lean_ctor_set(x_1200, 1, x_1198); +lean_ctor_set(x_1200, 2, x_1199); +lean_ctor_set(x_1181, 1, x_1200); +lean_ctor_set(x_1181, 0, x_1); +return x_1179; +} +} +else +{ +lean_object* x_1201; +lean_dec(x_1185); +lean_dec(x_1183); +lean_dec(x_1178); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1201 = lean_ctor_get(x_1186, 0); +lean_inc(x_1201); +lean_dec(x_1186); +lean_ctor_set(x_1181, 0, x_1201); +return x_1179; +} +} +else +{ +lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; +x_1202 = lean_ctor_get(x_1181, 0); +x_1203 = lean_ctor_get(x_1181, 1); +lean_inc(x_1203); +lean_inc(x_1202); +lean_dec(x_1181); +x_1204 = lean_ctor_get(x_1203, 2); +lean_inc(x_1204); +lean_inc(x_1204); +x_1205 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1204, x_1202); +if (lean_obj_tag(x_1205) == 0) +{ +size_t x_1206; size_t x_1207; uint8_t x_1208; +x_1206 = lean_ptr_addr(x_1178); +lean_dec(x_1178); +x_1207 = lean_ptr_addr(x_1202); +x_1208 = lean_usize_dec_eq(x_1206, x_1207); +if (x_1208 == 0) +{ +lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; +lean_dec(x_1); +x_1209 = lean_unsigned_to_nat(0u); +lean_inc(x_1202); +x_1210 = lean_array_set(x_3, x_1209, x_1202); +x_1211 = l_Lean_mkAppN(x_2, x_1210); +lean_dec(x_1210); +x_1212 = lean_ctor_get(x_1203, 0); +lean_inc(x_1212); +x_1213 = lean_ctor_get(x_1203, 1); +lean_inc(x_1213); +lean_dec(x_1203); +lean_inc(x_1211); +x_1214 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1204, x_1202, x_1211); +x_1215 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1215, 0, x_1212); +lean_ctor_set(x_1215, 1, x_1213); +lean_ctor_set(x_1215, 2, x_1214); +x_1216 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1216, 0, x_1211); +lean_ctor_set(x_1216, 1, x_1215); +lean_ctor_set(x_1179, 0, x_1216); +return x_1179; +} +else +{ +lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; +lean_dec(x_3); +lean_dec(x_2); +x_1217 = lean_ctor_get(x_1203, 0); +lean_inc(x_1217); +x_1218 = lean_ctor_get(x_1203, 1); +lean_inc(x_1218); +lean_dec(x_1203); +lean_inc(x_1); +x_1219 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1204, x_1202, x_1); +x_1220 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1220, 0, x_1217); +lean_ctor_set(x_1220, 1, x_1218); +lean_ctor_set(x_1220, 2, x_1219); +x_1221 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1221, 0, x_1); +lean_ctor_set(x_1221, 1, x_1220); +lean_ctor_set(x_1179, 0, x_1221); +return x_1179; +} +} +else +{ +lean_object* x_1222; lean_object* x_1223; +lean_dec(x_1204); +lean_dec(x_1202); +lean_dec(x_1178); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1222 = lean_ctor_get(x_1205, 0); +lean_inc(x_1222); +lean_dec(x_1205); +x_1223 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1223, 0, x_1222); +lean_ctor_set(x_1223, 1, x_1203); +lean_ctor_set(x_1179, 0, x_1223); +return x_1179; +} +} +} +else +{ +lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; +x_1224 = lean_ctor_get(x_1179, 0); +x_1225 = lean_ctor_get(x_1179, 1); +lean_inc(x_1225); +lean_inc(x_1224); +lean_dec(x_1179); +x_1226 = lean_ctor_get(x_1224, 0); +lean_inc(x_1226); +x_1227 = lean_ctor_get(x_1224, 1); +lean_inc(x_1227); +if (lean_is_exclusive(x_1224)) { + lean_ctor_release(x_1224, 0); + lean_ctor_release(x_1224, 1); + x_1228 = x_1224; +} else { + lean_dec_ref(x_1224); + x_1228 = lean_box(0); +} +x_1229 = lean_ctor_get(x_1227, 2); +lean_inc(x_1229); +lean_inc(x_1229); +x_1230 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1229, x_1226); +if (lean_obj_tag(x_1230) == 0) +{ +size_t x_1231; size_t x_1232; uint8_t x_1233; +x_1231 = lean_ptr_addr(x_1178); +lean_dec(x_1178); +x_1232 = lean_ptr_addr(x_1226); +x_1233 = lean_usize_dec_eq(x_1231, x_1232); +if (x_1233 == 0) +{ +lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; +lean_dec(x_1); +x_1234 = lean_unsigned_to_nat(0u); +lean_inc(x_1226); +x_1235 = lean_array_set(x_3, x_1234, x_1226); +x_1236 = l_Lean_mkAppN(x_2, x_1235); +lean_dec(x_1235); +x_1237 = lean_ctor_get(x_1227, 0); +lean_inc(x_1237); +x_1238 = lean_ctor_get(x_1227, 1); +lean_inc(x_1238); +lean_dec(x_1227); +lean_inc(x_1236); +x_1239 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1229, x_1226, x_1236); +x_1240 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1240, 0, x_1237); +lean_ctor_set(x_1240, 1, x_1238); +lean_ctor_set(x_1240, 2, x_1239); +if (lean_is_scalar(x_1228)) { + x_1241 = lean_alloc_ctor(0, 2, 0); +} else { + x_1241 = x_1228; +} +lean_ctor_set(x_1241, 0, x_1236); +lean_ctor_set(x_1241, 1, x_1240); +x_1242 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1242, 0, x_1241); +lean_ctor_set(x_1242, 1, x_1225); +return x_1242; +} +else +{ +lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; +lean_dec(x_3); +lean_dec(x_2); +x_1243 = lean_ctor_get(x_1227, 0); +lean_inc(x_1243); +x_1244 = lean_ctor_get(x_1227, 1); +lean_inc(x_1244); +lean_dec(x_1227); +lean_inc(x_1); +x_1245 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1229, x_1226, x_1); +x_1246 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1246, 0, x_1243); +lean_ctor_set(x_1246, 1, x_1244); +lean_ctor_set(x_1246, 2, x_1245); +if (lean_is_scalar(x_1228)) { + x_1247 = lean_alloc_ctor(0, 2, 0); +} else { + x_1247 = x_1228; +} +lean_ctor_set(x_1247, 0, x_1); +lean_ctor_set(x_1247, 1, x_1246); +x_1248 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1248, 0, x_1247); +lean_ctor_set(x_1248, 1, x_1225); +return x_1248; +} +} +else +{ +lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; +lean_dec(x_1229); +lean_dec(x_1226); +lean_dec(x_1178); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1249 = lean_ctor_get(x_1230, 0); +lean_inc(x_1249); +lean_dec(x_1230); +if (lean_is_scalar(x_1228)) { + x_1250 = lean_alloc_ctor(0, 2, 0); +} else { + x_1250 = x_1228; +} +lean_ctor_set(x_1250, 0, x_1249); +lean_ctor_set(x_1250, 1, x_1227); +x_1251 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1251, 0, x_1250); +lean_ctor_set(x_1251, 1, x_1225); +return x_1251; +} +} +} +else +{ +uint8_t x_1252; +lean_dec(x_1178); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1252 = !lean_is_exclusive(x_1179); +if (x_1252 == 0) +{ +return x_1179; +} +else +{ +lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; +x_1253 = lean_ctor_get(x_1179, 0); +x_1254 = lean_ctor_get(x_1179, 1); +lean_inc(x_1254); +lean_inc(x_1253); +lean_dec(x_1179); +x_1255 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1255, 0, x_1253); +lean_ctor_set(x_1255, 1, x_1254); +return x_1255; +} +} +} +block_1310: +{ +lean_object* x_1258; +lean_dec(x_1257); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_1258 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_1258) == 0) +{ +lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; uint8_t x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; +x_1259 = lean_ctor_get(x_1258, 0); +lean_inc(x_1259); +x_1260 = lean_ctor_get(x_1258, 1); +lean_inc(x_1260); +lean_dec(x_1258); +x_1261 = lean_ctor_get(x_1259, 0); +lean_inc(x_1261); +lean_dec(x_1259); +x_1262 = lean_array_get_size(x_3); +x_1263 = lean_unsigned_to_nat(0u); +x_1264 = lean_unsigned_to_nat(1u); +lean_inc(x_1262); +x_1265 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1265, 0, x_1263); +lean_ctor_set(x_1265, 1, x_1262); +lean_ctor_set(x_1265, 2, x_1264); +x_1266 = 0; +x_1267 = lean_box(x_1266); +lean_inc(x_3); +x_1268 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1268, 0, x_3); +lean_ctor_set(x_1268, 1, x_1267); +lean_inc(x_2); +x_1269 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_1261, x_1262, x_1265, x_1265, x_1268, x_1263, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_1260); +lean_dec(x_1265); +lean_dec(x_1262); +lean_dec(x_1261); +lean_dec(x_3); +if (lean_obj_tag(x_1269) == 0) +{ +lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; uint8_t x_1273; +x_1270 = lean_ctor_get(x_1269, 0); +lean_inc(x_1270); +x_1271 = lean_ctor_get(x_1270, 0); +lean_inc(x_1271); +x_1272 = lean_ctor_get(x_1271, 1); +lean_inc(x_1272); +x_1273 = lean_unbox(x_1272); +lean_dec(x_1272); +if (x_1273 == 0) +{ +uint8_t x_1274; +lean_dec(x_1271); +lean_dec(x_2); +x_1274 = !lean_is_exclusive(x_1269); +if (x_1274 == 0) +{ +lean_object* x_1275; uint8_t x_1276; +x_1275 = lean_ctor_get(x_1269, 0); +lean_dec(x_1275); +x_1276 = !lean_is_exclusive(x_1270); +if (x_1276 == 0) +{ +lean_object* x_1277; +x_1277 = lean_ctor_get(x_1270, 0); +lean_dec(x_1277); +lean_ctor_set(x_1270, 0, x_1); +return x_1269; +} +else +{ +lean_object* x_1278; lean_object* x_1279; +x_1278 = lean_ctor_get(x_1270, 1); +lean_inc(x_1278); +lean_dec(x_1270); +x_1279 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1279, 0, x_1); +lean_ctor_set(x_1279, 1, x_1278); +lean_ctor_set(x_1269, 0, x_1279); +return x_1269; +} +} +else +{ +lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; +x_1280 = lean_ctor_get(x_1269, 1); +lean_inc(x_1280); +lean_dec(x_1269); +x_1281 = lean_ctor_get(x_1270, 1); +lean_inc(x_1281); +if (lean_is_exclusive(x_1270)) { + lean_ctor_release(x_1270, 0); + lean_ctor_release(x_1270, 1); + x_1282 = x_1270; +} else { + lean_dec_ref(x_1270); + x_1282 = lean_box(0); +} +if (lean_is_scalar(x_1282)) { + x_1283 = lean_alloc_ctor(0, 2, 0); +} else { + x_1283 = x_1282; +} +lean_ctor_set(x_1283, 0, x_1); +lean_ctor_set(x_1283, 1, x_1281); +x_1284 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1284, 0, x_1283); +lean_ctor_set(x_1284, 1, x_1280); +return x_1284; +} +} +else +{ +uint8_t x_1285; +lean_dec(x_1); +x_1285 = !lean_is_exclusive(x_1269); +if (x_1285 == 0) +{ +lean_object* x_1286; uint8_t x_1287; +x_1286 = lean_ctor_get(x_1269, 0); +lean_dec(x_1286); +x_1287 = !lean_is_exclusive(x_1270); +if (x_1287 == 0) +{ +lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; +x_1288 = lean_ctor_get(x_1270, 0); +lean_dec(x_1288); +x_1289 = lean_ctor_get(x_1271, 0); +lean_inc(x_1289); +lean_dec(x_1271); +x_1290 = l_Lean_mkAppN(x_2, x_1289); +lean_dec(x_1289); +lean_ctor_set(x_1270, 0, x_1290); +return x_1269; +} +else +{ +lean_object* x_1291; lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; +x_1291 = lean_ctor_get(x_1270, 1); +lean_inc(x_1291); +lean_dec(x_1270); +x_1292 = lean_ctor_get(x_1271, 0); +lean_inc(x_1292); +lean_dec(x_1271); +x_1293 = l_Lean_mkAppN(x_2, x_1292); +lean_dec(x_1292); +x_1294 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1294, 0, x_1293); +lean_ctor_set(x_1294, 1, x_1291); +lean_ctor_set(x_1269, 0, x_1294); +return x_1269; +} +} +else +{ +lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; +x_1295 = lean_ctor_get(x_1269, 1); +lean_inc(x_1295); +lean_dec(x_1269); +x_1296 = lean_ctor_get(x_1270, 1); +lean_inc(x_1296); +if (lean_is_exclusive(x_1270)) { + lean_ctor_release(x_1270, 0); + lean_ctor_release(x_1270, 1); + x_1297 = x_1270; +} else { + lean_dec_ref(x_1270); + x_1297 = lean_box(0); +} +x_1298 = lean_ctor_get(x_1271, 0); +lean_inc(x_1298); +lean_dec(x_1271); +x_1299 = l_Lean_mkAppN(x_2, x_1298); +lean_dec(x_1298); +if (lean_is_scalar(x_1297)) { + x_1300 = lean_alloc_ctor(0, 2, 0); +} else { + x_1300 = x_1297; +} +lean_ctor_set(x_1300, 0, x_1299); +lean_ctor_set(x_1300, 1, x_1296); +x_1301 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1301, 0, x_1300); +lean_ctor_set(x_1301, 1, x_1295); +return x_1301; +} +} +} +else +{ +uint8_t x_1302; +lean_dec(x_2); +lean_dec(x_1); +x_1302 = !lean_is_exclusive(x_1269); +if (x_1302 == 0) +{ +return x_1269; +} +else +{ +lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; +x_1303 = lean_ctor_get(x_1269, 0); +x_1304 = lean_ctor_get(x_1269, 1); +lean_inc(x_1304); +lean_inc(x_1303); +lean_dec(x_1269); +x_1305 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1305, 0, x_1303); +lean_ctor_set(x_1305, 1, x_1304); +return x_1305; +} +} +} +else +{ +uint8_t x_1306; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1306 = !lean_is_exclusive(x_1258); +if (x_1306 == 0) +{ +return x_1258; +} +else +{ +lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; +x_1307 = lean_ctor_get(x_1258, 0); +x_1308 = lean_ctor_get(x_1258, 1); +lean_inc(x_1308); +lean_inc(x_1307); +lean_dec(x_1258); +x_1309 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1309, 0, x_1307); +lean_ctor_set(x_1309, 1, x_1308); +return x_1309; +} +} +} +} +case 10: +{ +lean_object* x_1323; lean_object* x_1402; lean_object* x_1456; uint8_t x_1457; +lean_dec(x_4); +x_1456 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_1457 = l_Lean_Expr_isConstOf(x_2, x_1456); +if (x_1457 == 0) +{ +lean_object* x_1458; +x_1458 = lean_box(0); +x_1402 = x_1458; +goto block_1455; +} +else +{ +lean_object* x_1459; lean_object* x_1460; uint8_t x_1461; +x_1459 = lean_array_get_size(x_3); +x_1460 = lean_unsigned_to_nat(2u); +x_1461 = lean_nat_dec_eq(x_1459, x_1460); +if (x_1461 == 0) +{ +lean_object* x_1462; +lean_dec(x_1459); +x_1462 = lean_box(0); +x_1402 = x_1462; +goto block_1455; +} +else +{ +lean_object* x_1463; uint8_t x_1464; +x_1463 = lean_unsigned_to_nat(0u); +x_1464 = lean_nat_dec_lt(x_1463, x_1459); +lean_dec(x_1459); +if (x_1464 == 0) +{ +lean_object* x_1465; lean_object* x_1466; +x_1465 = l_Lean_instInhabitedExpr; +x_1466 = l_outOfBounds___rarg(x_1465); +x_1323 = x_1466; +goto block_1401; +} +else +{ +lean_object* x_1467; +x_1467 = lean_array_fget(x_3, x_1463); +x_1323 = x_1467; +goto block_1401; +} +} +} +block_1401: +{ +lean_object* x_1324; +lean_inc(x_1323); +x_1324 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_1323, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_1324) == 0) +{ +uint8_t x_1325; +x_1325 = !lean_is_exclusive(x_1324); +if (x_1325 == 0) +{ +lean_object* x_1326; uint8_t x_1327; +x_1326 = lean_ctor_get(x_1324, 0); +x_1327 = !lean_is_exclusive(x_1326); +if (x_1327 == 0) +{ +lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; +x_1328 = lean_ctor_get(x_1326, 0); +x_1329 = lean_ctor_get(x_1326, 1); +x_1330 = lean_ctor_get(x_1329, 2); +lean_inc(x_1330); +lean_inc(x_1330); +x_1331 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1330, x_1328); +if (lean_obj_tag(x_1331) == 0) +{ +size_t x_1332; size_t x_1333; uint8_t x_1334; +x_1332 = lean_ptr_addr(x_1323); +lean_dec(x_1323); +x_1333 = lean_ptr_addr(x_1328); +x_1334 = lean_usize_dec_eq(x_1332, x_1333); +if (x_1334 == 0) +{ +lean_object* x_1335; lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; +lean_dec(x_1); +x_1335 = lean_unsigned_to_nat(0u); +lean_inc(x_1328); +x_1336 = lean_array_set(x_3, x_1335, x_1328); +x_1337 = l_Lean_mkAppN(x_2, x_1336); +lean_dec(x_1336); +x_1338 = lean_ctor_get(x_1329, 0); +lean_inc(x_1338); +x_1339 = lean_ctor_get(x_1329, 1); +lean_inc(x_1339); +lean_dec(x_1329); +lean_inc(x_1337); +x_1340 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1330, x_1328, x_1337); +x_1341 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1341, 0, x_1338); +lean_ctor_set(x_1341, 1, x_1339); +lean_ctor_set(x_1341, 2, x_1340); +lean_ctor_set(x_1326, 1, x_1341); +lean_ctor_set(x_1326, 0, x_1337); +return x_1324; +} +else +{ +lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; +lean_dec(x_3); +lean_dec(x_2); +x_1342 = lean_ctor_get(x_1329, 0); +lean_inc(x_1342); +x_1343 = lean_ctor_get(x_1329, 1); +lean_inc(x_1343); +lean_dec(x_1329); +lean_inc(x_1); +x_1344 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1330, x_1328, x_1); +x_1345 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1345, 0, x_1342); +lean_ctor_set(x_1345, 1, x_1343); +lean_ctor_set(x_1345, 2, x_1344); +lean_ctor_set(x_1326, 1, x_1345); +lean_ctor_set(x_1326, 0, x_1); +return x_1324; +} +} +else +{ +lean_object* x_1346; +lean_dec(x_1330); +lean_dec(x_1328); +lean_dec(x_1323); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1346 = lean_ctor_get(x_1331, 0); +lean_inc(x_1346); +lean_dec(x_1331); +lean_ctor_set(x_1326, 0, x_1346); +return x_1324; +} +} +else +{ +lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; +x_1347 = lean_ctor_get(x_1326, 0); +x_1348 = lean_ctor_get(x_1326, 1); +lean_inc(x_1348); +lean_inc(x_1347); +lean_dec(x_1326); +x_1349 = lean_ctor_get(x_1348, 2); +lean_inc(x_1349); +lean_inc(x_1349); +x_1350 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1349, x_1347); +if (lean_obj_tag(x_1350) == 0) +{ +size_t x_1351; size_t x_1352; uint8_t x_1353; +x_1351 = lean_ptr_addr(x_1323); +lean_dec(x_1323); +x_1352 = lean_ptr_addr(x_1347); +x_1353 = lean_usize_dec_eq(x_1351, x_1352); +if (x_1353 == 0) +{ +lean_object* x_1354; lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; +lean_dec(x_1); +x_1354 = lean_unsigned_to_nat(0u); +lean_inc(x_1347); +x_1355 = lean_array_set(x_3, x_1354, x_1347); +x_1356 = l_Lean_mkAppN(x_2, x_1355); +lean_dec(x_1355); +x_1357 = lean_ctor_get(x_1348, 0); +lean_inc(x_1357); +x_1358 = lean_ctor_get(x_1348, 1); +lean_inc(x_1358); +lean_dec(x_1348); +lean_inc(x_1356); +x_1359 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1349, x_1347, x_1356); +x_1360 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1360, 0, x_1357); +lean_ctor_set(x_1360, 1, x_1358); +lean_ctor_set(x_1360, 2, x_1359); +x_1361 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1361, 0, x_1356); +lean_ctor_set(x_1361, 1, x_1360); +lean_ctor_set(x_1324, 0, x_1361); +return x_1324; +} +else +{ +lean_object* x_1362; lean_object* x_1363; lean_object* x_1364; lean_object* x_1365; lean_object* x_1366; +lean_dec(x_3); +lean_dec(x_2); +x_1362 = lean_ctor_get(x_1348, 0); +lean_inc(x_1362); +x_1363 = lean_ctor_get(x_1348, 1); +lean_inc(x_1363); +lean_dec(x_1348); +lean_inc(x_1); +x_1364 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1349, x_1347, x_1); +x_1365 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1365, 0, x_1362); +lean_ctor_set(x_1365, 1, x_1363); +lean_ctor_set(x_1365, 2, x_1364); +x_1366 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1366, 0, x_1); +lean_ctor_set(x_1366, 1, x_1365); +lean_ctor_set(x_1324, 0, x_1366); +return x_1324; +} +} +else +{ +lean_object* x_1367; lean_object* x_1368; +lean_dec(x_1349); +lean_dec(x_1347); +lean_dec(x_1323); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1367 = lean_ctor_get(x_1350, 0); +lean_inc(x_1367); +lean_dec(x_1350); +x_1368 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1368, 0, x_1367); +lean_ctor_set(x_1368, 1, x_1348); +lean_ctor_set(x_1324, 0, x_1368); +return x_1324; +} +} +} +else +{ +lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; +x_1369 = lean_ctor_get(x_1324, 0); +x_1370 = lean_ctor_get(x_1324, 1); +lean_inc(x_1370); +lean_inc(x_1369); +lean_dec(x_1324); +x_1371 = lean_ctor_get(x_1369, 0); +lean_inc(x_1371); +x_1372 = lean_ctor_get(x_1369, 1); +lean_inc(x_1372); +if (lean_is_exclusive(x_1369)) { + lean_ctor_release(x_1369, 0); + lean_ctor_release(x_1369, 1); + x_1373 = x_1369; +} else { + lean_dec_ref(x_1369); + x_1373 = lean_box(0); +} +x_1374 = lean_ctor_get(x_1372, 2); +lean_inc(x_1374); +lean_inc(x_1374); +x_1375 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1374, x_1371); +if (lean_obj_tag(x_1375) == 0) +{ +size_t x_1376; size_t x_1377; uint8_t x_1378; +x_1376 = lean_ptr_addr(x_1323); +lean_dec(x_1323); +x_1377 = lean_ptr_addr(x_1371); +x_1378 = lean_usize_dec_eq(x_1376, x_1377); +if (x_1378 == 0) +{ +lean_object* x_1379; lean_object* x_1380; lean_object* x_1381; lean_object* x_1382; lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; lean_object* x_1387; +lean_dec(x_1); +x_1379 = lean_unsigned_to_nat(0u); +lean_inc(x_1371); +x_1380 = lean_array_set(x_3, x_1379, x_1371); +x_1381 = l_Lean_mkAppN(x_2, x_1380); +lean_dec(x_1380); +x_1382 = lean_ctor_get(x_1372, 0); +lean_inc(x_1382); +x_1383 = lean_ctor_get(x_1372, 1); +lean_inc(x_1383); +lean_dec(x_1372); +lean_inc(x_1381); +x_1384 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1374, x_1371, x_1381); +x_1385 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1385, 0, x_1382); +lean_ctor_set(x_1385, 1, x_1383); +lean_ctor_set(x_1385, 2, x_1384); +if (lean_is_scalar(x_1373)) { + x_1386 = lean_alloc_ctor(0, 2, 0); +} else { + x_1386 = x_1373; +} +lean_ctor_set(x_1386, 0, x_1381); +lean_ctor_set(x_1386, 1, x_1385); +x_1387 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1387, 0, x_1386); +lean_ctor_set(x_1387, 1, x_1370); +return x_1387; +} +else +{ +lean_object* x_1388; lean_object* x_1389; lean_object* x_1390; lean_object* x_1391; lean_object* x_1392; lean_object* x_1393; +lean_dec(x_3); +lean_dec(x_2); +x_1388 = lean_ctor_get(x_1372, 0); +lean_inc(x_1388); +x_1389 = lean_ctor_get(x_1372, 1); +lean_inc(x_1389); +lean_dec(x_1372); +lean_inc(x_1); +x_1390 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1374, x_1371, x_1); +x_1391 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1391, 0, x_1388); +lean_ctor_set(x_1391, 1, x_1389); +lean_ctor_set(x_1391, 2, x_1390); +if (lean_is_scalar(x_1373)) { + x_1392 = lean_alloc_ctor(0, 2, 0); +} else { + x_1392 = x_1373; +} +lean_ctor_set(x_1392, 0, x_1); +lean_ctor_set(x_1392, 1, x_1391); +x_1393 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1393, 0, x_1392); +lean_ctor_set(x_1393, 1, x_1370); +return x_1393; +} +} +else +{ +lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; +lean_dec(x_1374); +lean_dec(x_1371); +lean_dec(x_1323); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1394 = lean_ctor_get(x_1375, 0); +lean_inc(x_1394); +lean_dec(x_1375); +if (lean_is_scalar(x_1373)) { + x_1395 = lean_alloc_ctor(0, 2, 0); +} else { + x_1395 = x_1373; +} +lean_ctor_set(x_1395, 0, x_1394); +lean_ctor_set(x_1395, 1, x_1372); +x_1396 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1396, 0, x_1395); +lean_ctor_set(x_1396, 1, x_1370); +return x_1396; +} +} +} +else +{ +uint8_t x_1397; +lean_dec(x_1323); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1397 = !lean_is_exclusive(x_1324); +if (x_1397 == 0) +{ +return x_1324; +} +else +{ +lean_object* x_1398; lean_object* x_1399; lean_object* x_1400; +x_1398 = lean_ctor_get(x_1324, 0); +x_1399 = lean_ctor_get(x_1324, 1); +lean_inc(x_1399); +lean_inc(x_1398); +lean_dec(x_1324); +x_1400 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1400, 0, x_1398); +lean_ctor_set(x_1400, 1, x_1399); +return x_1400; +} +} +} +block_1455: +{ +lean_object* x_1403; +lean_dec(x_1402); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_1403 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_1403) == 0) +{ +lean_object* x_1404; lean_object* x_1405; lean_object* x_1406; lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; uint8_t x_1411; lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; +x_1404 = lean_ctor_get(x_1403, 0); +lean_inc(x_1404); +x_1405 = lean_ctor_get(x_1403, 1); +lean_inc(x_1405); +lean_dec(x_1403); +x_1406 = lean_ctor_get(x_1404, 0); +lean_inc(x_1406); +lean_dec(x_1404); +x_1407 = lean_array_get_size(x_3); +x_1408 = lean_unsigned_to_nat(0u); +x_1409 = lean_unsigned_to_nat(1u); +lean_inc(x_1407); +x_1410 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1410, 0, x_1408); +lean_ctor_set(x_1410, 1, x_1407); +lean_ctor_set(x_1410, 2, x_1409); +x_1411 = 0; +x_1412 = lean_box(x_1411); +lean_inc(x_3); +x_1413 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1413, 0, x_3); +lean_ctor_set(x_1413, 1, x_1412); +lean_inc(x_2); +x_1414 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_1406, x_1407, x_1410, x_1410, x_1413, x_1408, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_1405); +lean_dec(x_1410); +lean_dec(x_1407); +lean_dec(x_1406); +lean_dec(x_3); +if (lean_obj_tag(x_1414) == 0) +{ +lean_object* x_1415; lean_object* x_1416; lean_object* x_1417; uint8_t x_1418; +x_1415 = lean_ctor_get(x_1414, 0); +lean_inc(x_1415); +x_1416 = lean_ctor_get(x_1415, 0); +lean_inc(x_1416); +x_1417 = lean_ctor_get(x_1416, 1); +lean_inc(x_1417); +x_1418 = lean_unbox(x_1417); +lean_dec(x_1417); +if (x_1418 == 0) +{ +uint8_t x_1419; +lean_dec(x_1416); +lean_dec(x_2); +x_1419 = !lean_is_exclusive(x_1414); +if (x_1419 == 0) +{ +lean_object* x_1420; uint8_t x_1421; +x_1420 = lean_ctor_get(x_1414, 0); +lean_dec(x_1420); +x_1421 = !lean_is_exclusive(x_1415); +if (x_1421 == 0) +{ +lean_object* x_1422; +x_1422 = lean_ctor_get(x_1415, 0); +lean_dec(x_1422); +lean_ctor_set(x_1415, 0, x_1); +return x_1414; +} +else +{ +lean_object* x_1423; lean_object* x_1424; +x_1423 = lean_ctor_get(x_1415, 1); +lean_inc(x_1423); +lean_dec(x_1415); +x_1424 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1424, 0, x_1); +lean_ctor_set(x_1424, 1, x_1423); +lean_ctor_set(x_1414, 0, x_1424); +return x_1414; +} +} +else +{ +lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; lean_object* x_1428; lean_object* x_1429; +x_1425 = lean_ctor_get(x_1414, 1); +lean_inc(x_1425); +lean_dec(x_1414); +x_1426 = lean_ctor_get(x_1415, 1); +lean_inc(x_1426); +if (lean_is_exclusive(x_1415)) { + lean_ctor_release(x_1415, 0); + lean_ctor_release(x_1415, 1); + x_1427 = x_1415; +} else { + lean_dec_ref(x_1415); + x_1427 = lean_box(0); +} +if (lean_is_scalar(x_1427)) { + x_1428 = lean_alloc_ctor(0, 2, 0); +} else { + x_1428 = x_1427; +} +lean_ctor_set(x_1428, 0, x_1); +lean_ctor_set(x_1428, 1, x_1426); +x_1429 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1429, 0, x_1428); +lean_ctor_set(x_1429, 1, x_1425); +return x_1429; +} +} +else +{ +uint8_t x_1430; +lean_dec(x_1); +x_1430 = !lean_is_exclusive(x_1414); +if (x_1430 == 0) +{ +lean_object* x_1431; uint8_t x_1432; +x_1431 = lean_ctor_get(x_1414, 0); +lean_dec(x_1431); +x_1432 = !lean_is_exclusive(x_1415); +if (x_1432 == 0) +{ +lean_object* x_1433; lean_object* x_1434; lean_object* x_1435; +x_1433 = lean_ctor_get(x_1415, 0); +lean_dec(x_1433); +x_1434 = lean_ctor_get(x_1416, 0); +lean_inc(x_1434); +lean_dec(x_1416); +x_1435 = l_Lean_mkAppN(x_2, x_1434); +lean_dec(x_1434); +lean_ctor_set(x_1415, 0, x_1435); +return x_1414; +} +else +{ +lean_object* x_1436; lean_object* x_1437; lean_object* x_1438; lean_object* x_1439; +x_1436 = lean_ctor_get(x_1415, 1); +lean_inc(x_1436); +lean_dec(x_1415); +x_1437 = lean_ctor_get(x_1416, 0); +lean_inc(x_1437); +lean_dec(x_1416); +x_1438 = l_Lean_mkAppN(x_2, x_1437); +lean_dec(x_1437); +x_1439 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1439, 0, x_1438); +lean_ctor_set(x_1439, 1, x_1436); +lean_ctor_set(x_1414, 0, x_1439); +return x_1414; +} +} +else +{ +lean_object* x_1440; lean_object* x_1441; lean_object* x_1442; lean_object* x_1443; lean_object* x_1444; lean_object* x_1445; lean_object* x_1446; +x_1440 = lean_ctor_get(x_1414, 1); +lean_inc(x_1440); +lean_dec(x_1414); +x_1441 = lean_ctor_get(x_1415, 1); +lean_inc(x_1441); +if (lean_is_exclusive(x_1415)) { + lean_ctor_release(x_1415, 0); + lean_ctor_release(x_1415, 1); + x_1442 = x_1415; +} else { + lean_dec_ref(x_1415); + x_1442 = lean_box(0); +} +x_1443 = lean_ctor_get(x_1416, 0); +lean_inc(x_1443); +lean_dec(x_1416); +x_1444 = l_Lean_mkAppN(x_2, x_1443); +lean_dec(x_1443); +if (lean_is_scalar(x_1442)) { + x_1445 = lean_alloc_ctor(0, 2, 0); +} else { + x_1445 = x_1442; +} +lean_ctor_set(x_1445, 0, x_1444); +lean_ctor_set(x_1445, 1, x_1441); +x_1446 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1446, 0, x_1445); +lean_ctor_set(x_1446, 1, x_1440); +return x_1446; +} +} +} +else +{ +uint8_t x_1447; +lean_dec(x_2); +lean_dec(x_1); +x_1447 = !lean_is_exclusive(x_1414); +if (x_1447 == 0) +{ +return x_1414; +} +else +{ +lean_object* x_1448; lean_object* x_1449; lean_object* x_1450; +x_1448 = lean_ctor_get(x_1414, 0); +x_1449 = lean_ctor_get(x_1414, 1); +lean_inc(x_1449); +lean_inc(x_1448); +lean_dec(x_1414); +x_1450 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1450, 0, x_1448); +lean_ctor_set(x_1450, 1, x_1449); +return x_1450; +} +} +} +else +{ +uint8_t x_1451; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1451 = !lean_is_exclusive(x_1403); +if (x_1451 == 0) +{ +return x_1403; +} +else +{ +lean_object* x_1452; lean_object* x_1453; lean_object* x_1454; +x_1452 = lean_ctor_get(x_1403, 0); +x_1453 = lean_ctor_get(x_1403, 1); +lean_inc(x_1453); +lean_inc(x_1452); +lean_dec(x_1403); +x_1454 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1454, 0, x_1452); +lean_ctor_set(x_1454, 1, x_1453); +return x_1454; +} +} +} +} +default: +{ +lean_object* x_1468; lean_object* x_1547; lean_object* x_1601; uint8_t x_1602; +lean_dec(x_4); +x_1601 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; +x_1602 = l_Lean_Expr_isConstOf(x_2, x_1601); +if (x_1602 == 0) +{ +lean_object* x_1603; +x_1603 = lean_box(0); +x_1547 = x_1603; +goto block_1600; +} +else +{ +lean_object* x_1604; lean_object* x_1605; uint8_t x_1606; +x_1604 = lean_array_get_size(x_3); +x_1605 = lean_unsigned_to_nat(2u); +x_1606 = lean_nat_dec_eq(x_1604, x_1605); +if (x_1606 == 0) +{ +lean_object* x_1607; +lean_dec(x_1604); +x_1607 = lean_box(0); +x_1547 = x_1607; +goto block_1600; +} +else +{ +lean_object* x_1608; uint8_t x_1609; +x_1608 = lean_unsigned_to_nat(0u); +x_1609 = lean_nat_dec_lt(x_1608, x_1604); +lean_dec(x_1604); +if (x_1609 == 0) +{ +lean_object* x_1610; lean_object* x_1611; +x_1610 = l_Lean_instInhabitedExpr; +x_1611 = l_outOfBounds___rarg(x_1610); +x_1468 = x_1611; +goto block_1546; +} +else +{ +lean_object* x_1612; +x_1612 = lean_array_fget(x_3, x_1608); +x_1468 = x_1612; +goto block_1546; +} +} +} +block_1546: +{ +lean_object* x_1469; +lean_inc(x_1468); +x_1469 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_1468, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_1469) == 0) +{ +uint8_t x_1470; +x_1470 = !lean_is_exclusive(x_1469); +if (x_1470 == 0) +{ +lean_object* x_1471; uint8_t x_1472; +x_1471 = lean_ctor_get(x_1469, 0); +x_1472 = !lean_is_exclusive(x_1471); +if (x_1472 == 0) +{ +lean_object* x_1473; lean_object* x_1474; lean_object* x_1475; lean_object* x_1476; +x_1473 = lean_ctor_get(x_1471, 0); +x_1474 = lean_ctor_get(x_1471, 1); +x_1475 = lean_ctor_get(x_1474, 2); +lean_inc(x_1475); +lean_inc(x_1475); +x_1476 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1475, x_1473); +if (lean_obj_tag(x_1476) == 0) +{ +size_t x_1477; size_t x_1478; uint8_t x_1479; +x_1477 = lean_ptr_addr(x_1468); +lean_dec(x_1468); +x_1478 = lean_ptr_addr(x_1473); +x_1479 = lean_usize_dec_eq(x_1477, x_1478); +if (x_1479 == 0) +{ +lean_object* x_1480; lean_object* x_1481; lean_object* x_1482; lean_object* x_1483; lean_object* x_1484; lean_object* x_1485; lean_object* x_1486; +lean_dec(x_1); +x_1480 = lean_unsigned_to_nat(0u); +lean_inc(x_1473); +x_1481 = lean_array_set(x_3, x_1480, x_1473); +x_1482 = l_Lean_mkAppN(x_2, x_1481); +lean_dec(x_1481); +x_1483 = lean_ctor_get(x_1474, 0); +lean_inc(x_1483); +x_1484 = lean_ctor_get(x_1474, 1); +lean_inc(x_1484); +lean_dec(x_1474); +lean_inc(x_1482); +x_1485 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1475, x_1473, x_1482); +x_1486 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1486, 0, x_1483); +lean_ctor_set(x_1486, 1, x_1484); +lean_ctor_set(x_1486, 2, x_1485); +lean_ctor_set(x_1471, 1, x_1486); +lean_ctor_set(x_1471, 0, x_1482); +return x_1469; +} +else +{ +lean_object* x_1487; lean_object* x_1488; lean_object* x_1489; lean_object* x_1490; +lean_dec(x_3); +lean_dec(x_2); +x_1487 = lean_ctor_get(x_1474, 0); +lean_inc(x_1487); +x_1488 = lean_ctor_get(x_1474, 1); +lean_inc(x_1488); +lean_dec(x_1474); +lean_inc(x_1); +x_1489 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1475, x_1473, x_1); +x_1490 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1490, 0, x_1487); +lean_ctor_set(x_1490, 1, x_1488); +lean_ctor_set(x_1490, 2, x_1489); +lean_ctor_set(x_1471, 1, x_1490); +lean_ctor_set(x_1471, 0, x_1); +return x_1469; +} +} +else +{ +lean_object* x_1491; +lean_dec(x_1475); +lean_dec(x_1473); +lean_dec(x_1468); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1491 = lean_ctor_get(x_1476, 0); +lean_inc(x_1491); +lean_dec(x_1476); +lean_ctor_set(x_1471, 0, x_1491); +return x_1469; +} +} +else +{ +lean_object* x_1492; lean_object* x_1493; lean_object* x_1494; lean_object* x_1495; +x_1492 = lean_ctor_get(x_1471, 0); +x_1493 = lean_ctor_get(x_1471, 1); +lean_inc(x_1493); +lean_inc(x_1492); +lean_dec(x_1471); +x_1494 = lean_ctor_get(x_1493, 2); +lean_inc(x_1494); +lean_inc(x_1494); +x_1495 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1494, x_1492); +if (lean_obj_tag(x_1495) == 0) +{ +size_t x_1496; size_t x_1497; uint8_t x_1498; +x_1496 = lean_ptr_addr(x_1468); +lean_dec(x_1468); +x_1497 = lean_ptr_addr(x_1492); +x_1498 = lean_usize_dec_eq(x_1496, x_1497); +if (x_1498 == 0) +{ +lean_object* x_1499; lean_object* x_1500; lean_object* x_1501; lean_object* x_1502; lean_object* x_1503; lean_object* x_1504; lean_object* x_1505; lean_object* x_1506; +lean_dec(x_1); +x_1499 = lean_unsigned_to_nat(0u); +lean_inc(x_1492); +x_1500 = lean_array_set(x_3, x_1499, x_1492); +x_1501 = l_Lean_mkAppN(x_2, x_1500); +lean_dec(x_1500); +x_1502 = lean_ctor_get(x_1493, 0); +lean_inc(x_1502); +x_1503 = lean_ctor_get(x_1493, 1); +lean_inc(x_1503); +lean_dec(x_1493); +lean_inc(x_1501); +x_1504 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1494, x_1492, x_1501); +x_1505 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1505, 0, x_1502); +lean_ctor_set(x_1505, 1, x_1503); +lean_ctor_set(x_1505, 2, x_1504); +x_1506 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1506, 0, x_1501); +lean_ctor_set(x_1506, 1, x_1505); +lean_ctor_set(x_1469, 0, x_1506); +return x_1469; +} +else +{ +lean_object* x_1507; lean_object* x_1508; lean_object* x_1509; lean_object* x_1510; lean_object* x_1511; +lean_dec(x_3); +lean_dec(x_2); +x_1507 = lean_ctor_get(x_1493, 0); +lean_inc(x_1507); +x_1508 = lean_ctor_get(x_1493, 1); +lean_inc(x_1508); +lean_dec(x_1493); +lean_inc(x_1); +x_1509 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1494, x_1492, x_1); +x_1510 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1510, 0, x_1507); +lean_ctor_set(x_1510, 1, x_1508); +lean_ctor_set(x_1510, 2, x_1509); +x_1511 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1511, 0, x_1); +lean_ctor_set(x_1511, 1, x_1510); +lean_ctor_set(x_1469, 0, x_1511); +return x_1469; +} +} +else +{ +lean_object* x_1512; lean_object* x_1513; +lean_dec(x_1494); +lean_dec(x_1492); +lean_dec(x_1468); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1512 = lean_ctor_get(x_1495, 0); +lean_inc(x_1512); +lean_dec(x_1495); +x_1513 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1513, 0, x_1512); +lean_ctor_set(x_1513, 1, x_1493); +lean_ctor_set(x_1469, 0, x_1513); +return x_1469; +} +} +} +else +{ +lean_object* x_1514; lean_object* x_1515; lean_object* x_1516; lean_object* x_1517; lean_object* x_1518; lean_object* x_1519; lean_object* x_1520; +x_1514 = lean_ctor_get(x_1469, 0); +x_1515 = lean_ctor_get(x_1469, 1); +lean_inc(x_1515); +lean_inc(x_1514); +lean_dec(x_1469); +x_1516 = lean_ctor_get(x_1514, 0); +lean_inc(x_1516); +x_1517 = lean_ctor_get(x_1514, 1); +lean_inc(x_1517); +if (lean_is_exclusive(x_1514)) { + lean_ctor_release(x_1514, 0); + lean_ctor_release(x_1514, 1); + x_1518 = x_1514; +} else { + lean_dec_ref(x_1514); + x_1518 = lean_box(0); +} +x_1519 = lean_ctor_get(x_1517, 2); +lean_inc(x_1519); +lean_inc(x_1519); +x_1520 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_1519, x_1516); +if (lean_obj_tag(x_1520) == 0) +{ +size_t x_1521; size_t x_1522; uint8_t x_1523; +x_1521 = lean_ptr_addr(x_1468); +lean_dec(x_1468); +x_1522 = lean_ptr_addr(x_1516); +x_1523 = lean_usize_dec_eq(x_1521, x_1522); +if (x_1523 == 0) +{ +lean_object* x_1524; lean_object* x_1525; lean_object* x_1526; lean_object* x_1527; lean_object* x_1528; lean_object* x_1529; lean_object* x_1530; lean_object* x_1531; lean_object* x_1532; +lean_dec(x_1); +x_1524 = lean_unsigned_to_nat(0u); +lean_inc(x_1516); +x_1525 = lean_array_set(x_3, x_1524, x_1516); +x_1526 = l_Lean_mkAppN(x_2, x_1525); +lean_dec(x_1525); +x_1527 = lean_ctor_get(x_1517, 0); +lean_inc(x_1527); +x_1528 = lean_ctor_get(x_1517, 1); +lean_inc(x_1528); +lean_dec(x_1517); +lean_inc(x_1526); +x_1529 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1519, x_1516, x_1526); +x_1530 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1530, 0, x_1527); +lean_ctor_set(x_1530, 1, x_1528); +lean_ctor_set(x_1530, 2, x_1529); +if (lean_is_scalar(x_1518)) { + x_1531 = lean_alloc_ctor(0, 2, 0); +} else { + x_1531 = x_1518; +} +lean_ctor_set(x_1531, 0, x_1526); +lean_ctor_set(x_1531, 1, x_1530); +x_1532 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1532, 0, x_1531); +lean_ctor_set(x_1532, 1, x_1515); +return x_1532; +} +else +{ +lean_object* x_1533; lean_object* x_1534; lean_object* x_1535; lean_object* x_1536; lean_object* x_1537; lean_object* x_1538; +lean_dec(x_3); +lean_dec(x_2); +x_1533 = lean_ctor_get(x_1517, 0); +lean_inc(x_1533); +x_1534 = lean_ctor_get(x_1517, 1); +lean_inc(x_1534); +lean_dec(x_1517); +lean_inc(x_1); +x_1535 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_1519, x_1516, x_1); +x_1536 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1536, 0, x_1533); +lean_ctor_set(x_1536, 1, x_1534); +lean_ctor_set(x_1536, 2, x_1535); +if (lean_is_scalar(x_1518)) { + x_1537 = lean_alloc_ctor(0, 2, 0); +} else { + x_1537 = x_1518; +} +lean_ctor_set(x_1537, 0, x_1); +lean_ctor_set(x_1537, 1, x_1536); +x_1538 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1538, 0, x_1537); +lean_ctor_set(x_1538, 1, x_1515); +return x_1538; +} +} +else +{ +lean_object* x_1539; lean_object* x_1540; lean_object* x_1541; +lean_dec(x_1519); +lean_dec(x_1516); +lean_dec(x_1468); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1539 = lean_ctor_get(x_1520, 0); +lean_inc(x_1539); +lean_dec(x_1520); +if (lean_is_scalar(x_1518)) { + x_1540 = lean_alloc_ctor(0, 2, 0); +} else { + x_1540 = x_1518; +} +lean_ctor_set(x_1540, 0, x_1539); +lean_ctor_set(x_1540, 1, x_1517); +x_1541 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1541, 0, x_1540); +lean_ctor_set(x_1541, 1, x_1515); +return x_1541; +} +} +} +else +{ +uint8_t x_1542; +lean_dec(x_1468); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1542 = !lean_is_exclusive(x_1469); +if (x_1542 == 0) +{ +return x_1469; +} +else +{ +lean_object* x_1543; lean_object* x_1544; lean_object* x_1545; +x_1543 = lean_ctor_get(x_1469, 0); +x_1544 = lean_ctor_get(x_1469, 1); +lean_inc(x_1544); +lean_inc(x_1543); +lean_dec(x_1469); +x_1545 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1545, 0, x_1543); +lean_ctor_set(x_1545, 1, x_1544); +return x_1545; +} +} +} +block_1600: +{ +lean_object* x_1548; +lean_dec(x_1547); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_1548 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_1548) == 0) +{ +lean_object* x_1549; lean_object* x_1550; lean_object* x_1551; lean_object* x_1552; lean_object* x_1553; lean_object* x_1554; lean_object* x_1555; uint8_t x_1556; lean_object* x_1557; lean_object* x_1558; lean_object* x_1559; +x_1549 = lean_ctor_get(x_1548, 0); +lean_inc(x_1549); +x_1550 = lean_ctor_get(x_1548, 1); +lean_inc(x_1550); +lean_dec(x_1548); +x_1551 = lean_ctor_get(x_1549, 0); +lean_inc(x_1551); +lean_dec(x_1549); +x_1552 = lean_array_get_size(x_3); +x_1553 = lean_unsigned_to_nat(0u); +x_1554 = lean_unsigned_to_nat(1u); +lean_inc(x_1552); +x_1555 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_1555, 0, x_1553); +lean_ctor_set(x_1555, 1, x_1552); +lean_ctor_set(x_1555, 2, x_1554); +x_1556 = 0; +x_1557 = lean_box(x_1556); +lean_inc(x_3); +x_1558 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1558, 0, x_3); +lean_ctor_set(x_1558, 1, x_1557); +lean_inc(x_2); +x_1559 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_3, x_1551, x_1552, x_1555, x_1555, x_1558, x_1553, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_1550); +lean_dec(x_1555); +lean_dec(x_1552); +lean_dec(x_1551); +lean_dec(x_3); +if (lean_obj_tag(x_1559) == 0) +{ +lean_object* x_1560; lean_object* x_1561; lean_object* x_1562; uint8_t x_1563; +x_1560 = lean_ctor_get(x_1559, 0); +lean_inc(x_1560); +x_1561 = lean_ctor_get(x_1560, 0); +lean_inc(x_1561); +x_1562 = lean_ctor_get(x_1561, 1); +lean_inc(x_1562); +x_1563 = lean_unbox(x_1562); +lean_dec(x_1562); +if (x_1563 == 0) +{ +uint8_t x_1564; +lean_dec(x_1561); +lean_dec(x_2); +x_1564 = !lean_is_exclusive(x_1559); +if (x_1564 == 0) +{ +lean_object* x_1565; uint8_t x_1566; +x_1565 = lean_ctor_get(x_1559, 0); +lean_dec(x_1565); +x_1566 = !lean_is_exclusive(x_1560); +if (x_1566 == 0) +{ +lean_object* x_1567; +x_1567 = lean_ctor_get(x_1560, 0); +lean_dec(x_1567); +lean_ctor_set(x_1560, 0, x_1); +return x_1559; +} +else +{ +lean_object* x_1568; lean_object* x_1569; +x_1568 = lean_ctor_get(x_1560, 1); +lean_inc(x_1568); +lean_dec(x_1560); +x_1569 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1569, 0, x_1); +lean_ctor_set(x_1569, 1, x_1568); +lean_ctor_set(x_1559, 0, x_1569); +return x_1559; +} +} +else +{ +lean_object* x_1570; lean_object* x_1571; lean_object* x_1572; lean_object* x_1573; lean_object* x_1574; +x_1570 = lean_ctor_get(x_1559, 1); +lean_inc(x_1570); +lean_dec(x_1559); +x_1571 = lean_ctor_get(x_1560, 1); +lean_inc(x_1571); +if (lean_is_exclusive(x_1560)) { + lean_ctor_release(x_1560, 0); + lean_ctor_release(x_1560, 1); + x_1572 = x_1560; +} else { + lean_dec_ref(x_1560); + x_1572 = lean_box(0); +} +if (lean_is_scalar(x_1572)) { + x_1573 = lean_alloc_ctor(0, 2, 0); +} else { + x_1573 = x_1572; +} +lean_ctor_set(x_1573, 0, x_1); +lean_ctor_set(x_1573, 1, x_1571); +x_1574 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1574, 0, x_1573); +lean_ctor_set(x_1574, 1, x_1570); +return x_1574; +} +} +else +{ +uint8_t x_1575; +lean_dec(x_1); +x_1575 = !lean_is_exclusive(x_1559); +if (x_1575 == 0) +{ +lean_object* x_1576; uint8_t x_1577; +x_1576 = lean_ctor_get(x_1559, 0); +lean_dec(x_1576); +x_1577 = !lean_is_exclusive(x_1560); +if (x_1577 == 0) +{ +lean_object* x_1578; lean_object* x_1579; lean_object* x_1580; +x_1578 = lean_ctor_get(x_1560, 0); +lean_dec(x_1578); +x_1579 = lean_ctor_get(x_1561, 0); +lean_inc(x_1579); +lean_dec(x_1561); +x_1580 = l_Lean_mkAppN(x_2, x_1579); +lean_dec(x_1579); +lean_ctor_set(x_1560, 0, x_1580); +return x_1559; +} +else +{ +lean_object* x_1581; lean_object* x_1582; lean_object* x_1583; lean_object* x_1584; +x_1581 = lean_ctor_get(x_1560, 1); +lean_inc(x_1581); +lean_dec(x_1560); +x_1582 = lean_ctor_get(x_1561, 0); +lean_inc(x_1582); +lean_dec(x_1561); +x_1583 = l_Lean_mkAppN(x_2, x_1582); +lean_dec(x_1582); +x_1584 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1584, 0, x_1583); +lean_ctor_set(x_1584, 1, x_1581); +lean_ctor_set(x_1559, 0, x_1584); +return x_1559; +} +} +else +{ +lean_object* x_1585; lean_object* x_1586; lean_object* x_1587; lean_object* x_1588; lean_object* x_1589; lean_object* x_1590; lean_object* x_1591; +x_1585 = lean_ctor_get(x_1559, 1); +lean_inc(x_1585); +lean_dec(x_1559); +x_1586 = lean_ctor_get(x_1560, 1); +lean_inc(x_1586); +if (lean_is_exclusive(x_1560)) { + lean_ctor_release(x_1560, 0); + lean_ctor_release(x_1560, 1); + x_1587 = x_1560; +} else { + lean_dec_ref(x_1560); + x_1587 = lean_box(0); +} +x_1588 = lean_ctor_get(x_1561, 0); +lean_inc(x_1588); +lean_dec(x_1561); +x_1589 = l_Lean_mkAppN(x_2, x_1588); +lean_dec(x_1588); +if (lean_is_scalar(x_1587)) { + x_1590 = lean_alloc_ctor(0, 2, 0); +} else { + x_1590 = x_1587; +} +lean_ctor_set(x_1590, 0, x_1589); +lean_ctor_set(x_1590, 1, x_1586); +x_1591 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1591, 0, x_1590); +lean_ctor_set(x_1591, 1, x_1585); +return x_1591; +} +} +} +else +{ +uint8_t x_1592; +lean_dec(x_2); +lean_dec(x_1); +x_1592 = !lean_is_exclusive(x_1559); +if (x_1592 == 0) +{ +return x_1559; +} +else +{ +lean_object* x_1593; lean_object* x_1594; lean_object* x_1595; +x_1593 = lean_ctor_get(x_1559, 0); +x_1594 = lean_ctor_get(x_1559, 1); +lean_inc(x_1594); +lean_inc(x_1593); +lean_dec(x_1559); +x_1595 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1595, 0, x_1593); +lean_ctor_set(x_1595, 1, x_1594); +return x_1595; +} +} +} +else +{ +uint8_t x_1596; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_1596 = !lean_is_exclusive(x_1548); +if (x_1596 == 0) +{ +return x_1548; +} +else +{ +lean_object* x_1597; lean_object* x_1598; lean_object* x_1599; +x_1597 = lean_ctor_get(x_1548, 0); +x_1598 = lean_ctor_get(x_1548, 1); +lean_inc(x_1598); +lean_inc(x_1597); +lean_dec(x_1548); +x_1599 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1599, 0, x_1597); +lean_ctor_set(x_1599, 1, x_1598); +return x_1599; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_ptr_addr(x_1); +x_9 = lean_usize_dec_eq(x_7, x_8); +if (x_9 == 0) +{ +x_2 = x_6; +goto _start; +} +else +{ +lean_object* x_11; +lean_inc(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_5); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_take(x_3, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; uint8_t x_34; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_12, 0); +x_16 = lean_ctor_get(x_12, 1); +x_17 = lean_array_get_size(x_16); +x_18 = lean_ptr_addr(x_1); +x_19 = lean_usize_to_uint64(x_18); +x_20 = 11; +x_21 = lean_uint64_mix_hash(x_19, x_20); +x_22 = 32; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = 16; +x_26 = lean_uint64_shift_right(x_24, x_25); +x_27 = lean_uint64_xor(x_24, x_26); +x_28 = lean_uint64_to_usize(x_27); +x_29 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_30 = 1; +x_31 = lean_usize_sub(x_29, x_30); +x_32 = lean_usize_land(x_28, x_31); +x_33 = lean_array_uget(x_16, x_32); +x_34 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_1, x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_35 = lean_unsigned_to_nat(1u); +x_36 = lean_nat_add(x_15, x_35); +lean_dec(x_15); +lean_inc(x_2); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_1); +lean_ctor_set(x_37, 1, x_2); +lean_ctor_set(x_37, 2, x_33); +x_38 = lean_array_uset(x_16, x_32, x_37); +x_39 = lean_unsigned_to_nat(4u); +x_40 = lean_nat_mul(x_36, x_39); +x_41 = lean_unsigned_to_nat(3u); +x_42 = lean_nat_div(x_40, x_41); +lean_dec(x_40); +x_43 = lean_array_get_size(x_38); +x_44 = lean_nat_dec_le(x_42, x_43); +lean_dec(x_43); +lean_dec(x_42); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_38); +lean_ctor_set(x_12, 1, x_45); +lean_ctor_set(x_12, 0, x_36); +x_46 = lean_st_ref_set(x_3, x_12, x_14); +x_47 = !lean_is_exclusive(x_46); +if (x_47 == 0) +{ +lean_object* x_48; +x_48 = lean_ctor_get(x_46, 0); +lean_dec(x_48); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_46, 0, x_10); +return x_46; +} +else +{ +lean_object* x_49; lean_object* x_50; +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +lean_dec(x_46); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_10); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +else +{ +lean_object* x_51; uint8_t x_52; +lean_ctor_set(x_12, 1, x_38); +lean_ctor_set(x_12, 0, x_36); +x_51 = lean_st_ref_set(x_3, x_12, x_14); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) +{ +lean_object* x_53; +x_53 = lean_ctor_get(x_51, 0); +lean_dec(x_53); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_51, 0, x_10); +return x_51; +} +else +{ +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_10); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_56 = lean_box(0); +x_57 = lean_array_uset(x_16, x_32, x_56); +lean_inc(x_2); +x_58 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_33); +x_59 = lean_array_uset(x_57, x_32, x_58); +lean_ctor_set(x_12, 1, x_59); +x_60 = lean_st_ref_set(x_3, x_12, x_14); +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) +{ +lean_object* x_62; +x_62 = lean_ctor_get(x_60, 0); +lean_dec(x_62); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_60, 0, x_10); +return x_60; +} +else +{ +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_60, 1); +lean_inc(x_63); +lean_dec(x_60); +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_10); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; size_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; uint64_t x_76; uint64_t x_77; uint64_t x_78; size_t x_79; size_t x_80; size_t x_81; size_t x_82; size_t x_83; lean_object* x_84; uint8_t x_85; +x_65 = lean_ctor_get(x_10, 1); +x_66 = lean_ctor_get(x_12, 0); +x_67 = lean_ctor_get(x_12, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_12); +x_68 = lean_array_get_size(x_67); +x_69 = lean_ptr_addr(x_1); +x_70 = lean_usize_to_uint64(x_69); +x_71 = 11; +x_72 = lean_uint64_mix_hash(x_70, x_71); +x_73 = 32; +x_74 = lean_uint64_shift_right(x_72, x_73); +x_75 = lean_uint64_xor(x_72, x_74); +x_76 = 16; +x_77 = lean_uint64_shift_right(x_75, x_76); +x_78 = lean_uint64_xor(x_75, x_77); +x_79 = lean_uint64_to_usize(x_78); +x_80 = lean_usize_of_nat(x_68); +lean_dec(x_68); +x_81 = 1; +x_82 = lean_usize_sub(x_80, x_81); +x_83 = lean_usize_land(x_79, x_82); +x_84 = lean_array_uget(x_67, x_83); +x_85 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_1, x_84); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_86 = lean_unsigned_to_nat(1u); +x_87 = lean_nat_add(x_66, x_86); +lean_dec(x_66); +lean_inc(x_2); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_1); +lean_ctor_set(x_88, 1, x_2); +lean_ctor_set(x_88, 2, x_84); +x_89 = lean_array_uset(x_67, x_83, x_88); +x_90 = lean_unsigned_to_nat(4u); +x_91 = lean_nat_mul(x_87, x_90); +x_92 = lean_unsigned_to_nat(3u); +x_93 = lean_nat_div(x_91, x_92); +lean_dec(x_91); +x_94 = lean_array_get_size(x_89); +x_95 = lean_nat_dec_le(x_93, x_94); +lean_dec(x_94); +lean_dec(x_93); +if (x_95 == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_96 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_89); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_87); +lean_ctor_set(x_97, 1, x_96); +x_98 = lean_st_ref_set(x_3, x_97, x_65); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; +} else { + lean_dec_ref(x_98); + x_100 = lean_box(0); +} +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +if (lean_is_scalar(x_100)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_100; +} +lean_ctor_set(x_101, 0, x_10); +lean_ctor_set(x_101, 1, x_99); +return x_101; +} +else +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_87); +lean_ctor_set(x_102, 1, x_89); +x_103 = lean_st_ref_set(x_3, x_102, x_65); +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_105 = x_103; +} else { + lean_dec_ref(x_103); + x_105 = lean_box(0); +} +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +if (lean_is_scalar(x_105)) { + x_106 = lean_alloc_ctor(0, 2, 0); +} else { + x_106 = x_105; +} +lean_ctor_set(x_106, 0, x_10); +lean_ctor_set(x_106, 1, x_104); +return x_106; +} +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_107 = lean_box(0); +x_108 = lean_array_uset(x_67, x_83, x_107); +lean_inc(x_2); +x_109 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_84); +x_110 = lean_array_uset(x_108, x_83, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_66); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_st_ref_set(x_3, x_111, x_65); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_114 = x_112; +} else { + lean_dec_ref(x_112); + x_114 = lean_box(0); +} +lean_ctor_set(x_10, 1, x_4); +lean_ctor_set(x_10, 0, x_2); +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_114; +} +lean_ctor_set(x_115, 0, x_10); +lean_ctor_set(x_115, 1, x_113); +return x_115; +} +} +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; size_t x_122; uint64_t x_123; uint64_t x_124; uint64_t x_125; uint64_t x_126; uint64_t x_127; uint64_t x_128; uint64_t x_129; uint64_t x_130; uint64_t x_131; size_t x_132; size_t x_133; size_t x_134; size_t x_135; size_t x_136; lean_object* x_137; uint8_t x_138; +x_116 = lean_ctor_get(x_10, 0); +x_117 = lean_ctor_get(x_10, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_10); +x_118 = lean_ctor_get(x_116, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_116, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_120 = x_116; +} else { + lean_dec_ref(x_116); + x_120 = lean_box(0); +} +x_121 = lean_array_get_size(x_119); +x_122 = lean_ptr_addr(x_1); +x_123 = lean_usize_to_uint64(x_122); +x_124 = 11; +x_125 = lean_uint64_mix_hash(x_123, x_124); +x_126 = 32; +x_127 = lean_uint64_shift_right(x_125, x_126); +x_128 = lean_uint64_xor(x_125, x_127); +x_129 = 16; +x_130 = lean_uint64_shift_right(x_128, x_129); +x_131 = lean_uint64_xor(x_128, x_130); +x_132 = lean_uint64_to_usize(x_131); +x_133 = lean_usize_of_nat(x_121); +lean_dec(x_121); +x_134 = 1; +x_135 = lean_usize_sub(x_133, x_134); +x_136 = lean_usize_land(x_132, x_135); +x_137 = lean_array_uget(x_119, x_136); +x_138 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_1, x_137); +if (x_138 == 0) { -return x_677; -} -else +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; +x_139 = lean_unsigned_to_nat(1u); +x_140 = lean_nat_add(x_118, x_139); +lean_dec(x_118); +lean_inc(x_2); +x_141 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_141, 0, x_1); +lean_ctor_set(x_141, 1, x_2); +lean_ctor_set(x_141, 2, x_137); +x_142 = lean_array_uset(x_119, x_136, x_141); +x_143 = lean_unsigned_to_nat(4u); +x_144 = lean_nat_mul(x_140, x_143); +x_145 = lean_unsigned_to_nat(3u); +x_146 = lean_nat_div(x_144, x_145); +lean_dec(x_144); +x_147 = lean_array_get_size(x_142); +x_148 = lean_nat_dec_le(x_146, x_147); +lean_dec(x_147); +lean_dec(x_146); +if (x_148 == 0) { -lean_object* x_691; lean_object* x_692; lean_object* x_693; -x_691 = lean_ctor_get(x_677, 0); -x_692 = lean_ctor_get(x_677, 1); -lean_inc(x_692); -lean_inc(x_691); -lean_dec(x_677); -x_693 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_693, 0, x_691); -lean_ctor_set(x_693, 1, x_692); -return x_693; +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_149 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_142); +if (lean_is_scalar(x_120)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_120; } +lean_ctor_set(x_150, 0, x_140); +lean_ctor_set(x_150, 1, x_149); +x_151 = lean_st_ref_set(x_3, x_150, x_117); +x_152 = lean_ctor_get(x_151, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_151)) { + lean_ctor_release(x_151, 0); + lean_ctor_release(x_151, 1); + x_153 = x_151; +} else { + lean_dec_ref(x_151); + x_153 = lean_box(0); } +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_2); +lean_ctor_set(x_154, 1, x_4); +if (lean_is_scalar(x_153)) { + x_155 = lean_alloc_ctor(0, 2, 0); +} else { + x_155 = x_153; } -else -{ -uint8_t x_694; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_694 = !lean_is_exclusive(x_666); -if (x_694 == 0) -{ -return x_666; +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_152); +return x_155; } else { -lean_object* x_695; lean_object* x_696; lean_object* x_697; -x_695 = lean_ctor_get(x_666, 0); -x_696 = lean_ctor_get(x_666, 1); -lean_inc(x_696); -lean_inc(x_695); -lean_dec(x_666); -x_697 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_697, 0, x_695); -lean_ctor_set(x_697, 1, x_696); -return x_697; -} +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +if (lean_is_scalar(x_120)) { + x_156 = lean_alloc_ctor(0, 2, 0); +} else { + x_156 = x_120; } +lean_ctor_set(x_156, 0, x_140); +lean_ctor_set(x_156, 1, x_142); +x_157 = lean_st_ref_set(x_3, x_156, x_117); +x_158 = lean_ctor_get(x_157, 1); +lean_inc(x_158); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_159 = x_157; +} else { + lean_dec_ref(x_157); + x_159 = lean_box(0); } +x_160 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_160, 0, x_2); +lean_ctor_set(x_160, 1, x_4); +if (lean_is_scalar(x_159)) { + x_161 = lean_alloc_ctor(0, 2, 0); +} else { + x_161 = x_159; } -case 10: -{ -lean_object* x_711; lean_object* x_742; lean_object* x_776; uint8_t x_777; -lean_dec(x_4); -x_776 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_777 = l_Lean_Expr_isConstOf(x_2, x_776); -if (x_777 == 0) -{ -lean_object* x_778; -x_778 = lean_box(0); -x_742 = x_778; -goto block_775; +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_158); +return x_161; } -else -{ -lean_object* x_779; lean_object* x_780; uint8_t x_781; -x_779 = lean_array_get_size(x_3); -x_780 = lean_unsigned_to_nat(2u); -x_781 = lean_nat_dec_eq(x_779, x_780); -if (x_781 == 0) -{ -lean_object* x_782; -lean_dec(x_779); -x_782 = lean_box(0); -x_742 = x_782; -goto block_775; } else { -lean_object* x_783; uint8_t x_784; -x_783 = lean_unsigned_to_nat(0u); -x_784 = lean_nat_dec_lt(x_783, x_779); -lean_dec(x_779); -if (x_784 == 0) -{ -lean_object* x_785; lean_object* x_786; -x_785 = l_Lean_instInhabitedExpr; -x_786 = l_outOfBounds___rarg(x_785); -x_711 = x_786; -goto block_741; +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_162 = lean_box(0); +x_163 = lean_array_uset(x_119, x_136, x_162); +lean_inc(x_2); +x_164 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__6(x_1, x_2, x_137); +x_165 = lean_array_uset(x_163, x_136, x_164); +if (lean_is_scalar(x_120)) { + x_166 = lean_alloc_ctor(0, 2, 0); +} else { + x_166 = x_120; } -else -{ -lean_object* x_787; -x_787 = lean_array_fget(x_3, x_783); -x_711 = x_787; -goto block_741; +lean_ctor_set(x_166, 0, x_118); +lean_ctor_set(x_166, 1, x_165); +x_167 = lean_st_ref_set(x_3, x_166, x_117); +x_168 = lean_ctor_get(x_167, 1); +lean_inc(x_168); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_169 = x_167; +} else { + lean_dec_ref(x_167); + x_169 = lean_box(0); } +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_2); +lean_ctor_set(x_170, 1, x_4); +if (lean_is_scalar(x_169)) { + x_171 = lean_alloc_ctor(0, 2, 0); +} else { + x_171 = x_169; } +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_168); +return x_171; } -block_741: -{ -lean_object* x_712; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_711); -x_712 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_711, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_712) == 0) -{ -lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; -x_713 = lean_ctor_get(x_712, 0); -lean_inc(x_713); -x_714 = lean_ctor_get(x_712, 1); -lean_inc(x_714); -lean_dec(x_712); -x_715 = lean_ctor_get(x_713, 0); -lean_inc(x_715); -x_716 = lean_ctor_get(x_713, 1); -lean_inc(x_716); -lean_dec(x_713); -x_717 = lean_ctor_get(x_716, 2); -lean_inc(x_717); -lean_inc(x_717); -x_718 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_717, x_715); -if (lean_obj_tag(x_718) == 0) -{ -size_t x_719; size_t x_720; uint8_t x_721; -x_719 = lean_ptr_addr(x_711); -lean_dec(x_711); -x_720 = lean_ptr_addr(x_715); -x_721 = lean_usize_dec_eq(x_719, x_720); -if (x_721 == 0) -{ -lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; -x_722 = lean_unsigned_to_nat(0u); -lean_inc(x_715); -x_723 = lean_array_set(x_3, x_722, x_715); -x_724 = l_Lean_mkAppN(x_2, x_723); -lean_dec(x_723); -x_725 = lean_ctor_get(x_716, 0); -lean_inc(x_725); -x_726 = lean_ctor_get(x_716, 1); -lean_inc(x_726); -lean_dec(x_716); -lean_inc(x_724); -x_727 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_717, x_715, x_724); -x_728 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_728, 0, x_725); -lean_ctor_set(x_728, 1, x_726); -lean_ctor_set(x_728, 2, x_727); -x_729 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_724, x_5, x_728, x_7, x_8, x_9, x_10, x_714); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_729; } -else -{ -lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; -lean_dec(x_3); -lean_dec(x_2); -x_730 = lean_ctor_get(x_716, 0); -lean_inc(x_730); -x_731 = lean_ctor_get(x_716, 1); -lean_inc(x_731); -lean_dec(x_716); -lean_inc(x_1); -x_732 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_717, x_715, x_1); -x_733 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_733, 0, x_730); -lean_ctor_set(x_733, 1, x_731); -lean_ctor_set(x_733, 2, x_732); -lean_inc(x_1); -x_734 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_733, x_7, x_8, x_9, x_10, x_714); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_734; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_735; lean_object* x_736; -lean_dec(x_717); -lean_dec(x_715); -lean_dec(x_711); -lean_dec(x_3); +lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_15 = l_Lean_Expr_forallE___override(x_1, x_2, x_3, x_4); +x_16 = lean_ptr_addr(x_2); lean_dec(x_2); -x_735 = lean_ctor_get(x_718, 0); -lean_inc(x_735); -lean_dec(x_718); -x_736 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_735, x_5, x_716, x_7, x_8, x_9, x_10, x_714); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_736; -} -} -else +x_17 = lean_ptr_addr(x_5); +x_18 = lean_usize_dec_eq(x_16, x_17); +if (x_18 == 0) { -uint8_t x_737; -lean_dec(x_711); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); +lean_object* x_19; lean_object* x_20; +lean_dec(x_15); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_737 = !lean_is_exclusive(x_712); -if (x_737 == 0) -{ -return x_712; +x_19 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_20 = lean_apply_8(x_6, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_20; } else { -lean_object* x_738; lean_object* x_739; lean_object* x_740; -x_738 = lean_ctor_get(x_712, 0); -x_739 = lean_ctor_get(x_712, 1); -lean_inc(x_739); -lean_inc(x_738); -lean_dec(x_712); -x_740 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_740, 0, x_738); -lean_ctor_set(x_740, 1, x_739); -return x_740; -} -} -} -block_775: -{ -lean_object* x_743; -lean_dec(x_742); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_743 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_743) == 0) -{ -lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; uint8_t x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; -x_744 = lean_ctor_get(x_743, 0); -lean_inc(x_744); -x_745 = lean_ctor_get(x_743, 1); -lean_inc(x_745); -lean_dec(x_743); -x_746 = lean_ctor_get(x_744, 0); -lean_inc(x_746); -lean_dec(x_744); -x_747 = lean_array_get_size(x_3); -x_748 = lean_unsigned_to_nat(0u); -x_749 = lean_unsigned_to_nat(1u); -x_750 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_750, 0, x_748); -lean_ctor_set(x_750, 1, x_747); -lean_ctor_set(x_750, 2, x_749); -x_751 = 0; -x_752 = lean_box(x_751); -x_753 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_753, 0, x_3); -lean_ctor_set(x_753, 1, x_752); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -lean_inc(x_2); -x_754 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_746, x_750, x_750, x_753, x_748, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_745); -lean_dec(x_750); -lean_dec(x_746); -if (lean_obj_tag(x_754) == 0) -{ -lean_object* x_755; lean_object* x_756; lean_object* x_757; uint8_t x_758; -x_755 = lean_ctor_get(x_754, 0); -lean_inc(x_755); -x_756 = lean_ctor_get(x_755, 0); -lean_inc(x_756); -x_757 = lean_ctor_get(x_756, 1); -lean_inc(x_757); -x_758 = lean_unbox(x_757); -lean_dec(x_757); -if (x_758 == 0) -{ -lean_object* x_759; lean_object* x_760; lean_object* x_761; -lean_dec(x_756); -lean_dec(x_2); -x_759 = lean_ctor_get(x_754, 1); -lean_inc(x_759); -lean_dec(x_754); -x_760 = lean_ctor_get(x_755, 1); -lean_inc(x_760); -lean_dec(x_755); -lean_inc(x_1); -x_761 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_760, x_7, x_8, x_9, x_10, x_759); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_761; -} -else +size_t x_21; size_t x_22; uint8_t x_23; +x_21 = lean_ptr_addr(x_3); +lean_dec(x_3); +x_22 = lean_ptr_addr(x_7); +x_23 = lean_usize_dec_eq(x_21, x_22); +if (x_23 == 0) { -lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; -x_762 = lean_ctor_get(x_754, 1); -lean_inc(x_762); -lean_dec(x_754); -x_763 = lean_ctor_get(x_755, 1); -lean_inc(x_763); -lean_dec(x_755); -x_764 = lean_ctor_get(x_756, 0); -lean_inc(x_764); -lean_dec(x_756); -x_765 = l_Lean_mkAppN(x_2, x_764); -lean_dec(x_764); -x_766 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_765, x_5, x_763, x_7, x_8, x_9, x_10, x_762); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_766; -} +lean_object* x_24; lean_object* x_25; +lean_dec(x_15); +x_24 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_25 = lean_apply_8(x_6, x_24, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_25; } else { -uint8_t x_767; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_767 = !lean_is_exclusive(x_754); -if (x_767 == 0) -{ -return x_754; -} -else +uint8_t x_26; +x_26 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(x_4, x_4); +if (x_26 == 0) { -lean_object* x_768; lean_object* x_769; lean_object* x_770; -x_768 = lean_ctor_get(x_754, 0); -x_769 = lean_ctor_get(x_754, 1); -lean_inc(x_769); -lean_inc(x_768); -lean_dec(x_754); -x_770 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_770, 0, x_768); -lean_ctor_set(x_770, 1, x_769); -return x_770; -} -} +lean_object* x_27; lean_object* x_28; +lean_dec(x_15); +x_27 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_28 = lean_apply_8(x_6, x_27, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_28; } else { -uint8_t x_771; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_29; lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_771 = !lean_is_exclusive(x_743); -if (x_771 == 0) -{ -return x_743; +x_29 = lean_apply_8(x_6, x_15, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_29; } -else -{ -lean_object* x_772; lean_object* x_773; lean_object* x_774; -x_772 = lean_ctor_get(x_743, 0); -x_773 = lean_ctor_get(x_743, 1); -lean_inc(x_773); -lean_inc(x_772); -lean_dec(x_743); -x_774 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_774, 0, x_772); -lean_ctor_set(x_774, 1, x_773); -return x_774; } } } } -default: -{ -lean_object* x_788; lean_object* x_819; lean_object* x_853; uint8_t x_854; -lean_dec(x_4); -x_853 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4; -x_854 = l_Lean_Expr_isConstOf(x_2, x_853); -if (x_854 == 0) +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1() { +_start: { -lean_object* x_855; -x_855 = lean_box(0); -x_819 = x_855; -goto block_852; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Canon", 28, 28); +return x_1; } -else -{ -lean_object* x_856; lean_object* x_857; uint8_t x_858; -x_856 = lean_array_get_size(x_3); -x_857 = lean_unsigned_to_nat(2u); -x_858 = lean_nat_dec_eq(x_856, x_857); -if (x_858 == 0) -{ -lean_object* x_859; -lean_dec(x_856); -x_859 = lean_box(0); -x_819 = x_859; -goto block_852; } -else +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2() { +_start: { -lean_object* x_860; uint8_t x_861; -x_860 = lean_unsigned_to_nat(0u); -x_861 = lean_nat_dec_lt(x_860, x_856); -lean_dec(x_856); -if (x_861 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Canon.canonImpl.visit", 37, 37); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3() { +_start: { -lean_object* x_862; lean_object* x_863; -x_862 = l_Lean_instInhabitedExpr; -x_863 = l_outOfBounds___rarg(x_862); -x_788 = x_863; -goto block_818; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4() { +_start: { -lean_object* x_864; -x_864 = lean_array_fget(x_3, x_860); -x_788 = x_864; -goto block_818; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1; +x_2 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2; +x_3 = lean_unsigned_to_nat(149u); +x_4 = lean_unsigned_to_nat(13u); +x_5 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } +static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} } -block_818: +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_789; -lean_inc(x_10); -lean_inc(x_9); +switch (lean_obj_tag(x_1)) { +case 5: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_unsigned_to_nat(0u); +x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_10); +x_12 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5; +lean_inc(x_11); +x_13 = lean_mk_array(x_11, x_12); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_11, x_14); +lean_dec(x_11); lean_inc(x_8); lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -lean_inc(x_788); -x_789 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_788, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_789) == 0) -{ -lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; -x_790 = lean_ctor_get(x_789, 0); -lean_inc(x_790); -x_791 = lean_ctor_get(x_789, 1); -lean_inc(x_791); -lean_dec(x_789); -x_792 = lean_ctor_get(x_790, 0); -lean_inc(x_792); -x_793 = lean_ctor_get(x_790, 1); -lean_inc(x_793); -lean_dec(x_790); -x_794 = lean_ctor_get(x_793, 2); -lean_inc(x_794); -lean_inc(x_794); -x_795 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_Canon_canonElemCore___spec__15(x_794, x_792); -if (lean_obj_tag(x_795) == 0) -{ -size_t x_796; size_t x_797; uint8_t x_798; -x_796 = lean_ptr_addr(x_788); -lean_dec(x_788); -x_797 = lean_ptr_addr(x_792); -x_798 = lean_usize_dec_eq(x_796, x_797); -if (x_798 == 0) -{ -lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; -x_799 = lean_unsigned_to_nat(0u); -lean_inc(x_792); -x_800 = lean_array_set(x_3, x_799, x_792); -x_801 = l_Lean_mkAppN(x_2, x_800); -lean_dec(x_800); -x_802 = lean_ctor_get(x_793, 0); -lean_inc(x_802); -x_803 = lean_ctor_get(x_793, 1); -lean_inc(x_803); -lean_dec(x_793); -lean_inc(x_801); -x_804 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_794, x_792, x_801); -x_805 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_805, 0, x_802); -lean_ctor_set(x_805, 1, x_803); -lean_ctor_set(x_805, 2, x_804); -x_806 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_801, x_5, x_805, x_7, x_8, x_9, x_10, x_791); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_806; -} -else +lean_inc(x_3); +lean_inc_n(x_1, 2); +x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9(x_1, x_1, x_13, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; -lean_dec(x_3); -lean_dec(x_2); -x_807 = lean_ctor_get(x_793, 0); -lean_inc(x_807); -x_808 = lean_ctor_get(x_793, 1); -lean_inc(x_808); -lean_dec(x_793); -lean_inc(x_1); -x_809 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_Canon_canonElemCore___spec__6(x_794, x_792, x_1); -x_810 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_810, 0, x_807); -lean_ctor_set(x_810, 1, x_808); -lean_ctor_set(x_810, 2, x_809); -lean_inc(x_1); -x_811 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_810, x_7, x_8, x_9, x_10, x_791); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_19, x_3, x_20, x_5, x_6, x_7, x_8, x_18); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -return x_811; -} -} -else -{ -lean_object* x_812; lean_object* x_813; -lean_dec(x_794); -lean_dec(x_792); -lean_dec(x_788); lean_dec(x_3); -lean_dec(x_2); -x_812 = lean_ctor_get(x_795, 0); -lean_inc(x_812); -lean_dec(x_795); -x_813 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_812, x_5, x_793, x_7, x_8, x_9, x_10, x_791); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_813; -} +return x_21; } else { -uint8_t x_814; -lean_dec(x_788); -lean_dec(x_10); -lean_dec(x_9); +uint8_t x_22; lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_814 = !lean_is_exclusive(x_789); -if (x_814 == 0) +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) { -return x_789; +return x_16; } else { -lean_object* x_815; lean_object* x_816; lean_object* x_817; -x_815 = lean_ctor_get(x_789, 0); -x_816 = lean_ctor_get(x_789, 1); -lean_inc(x_816); -lean_inc(x_815); -lean_dec(x_789); -x_817 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_817, 0, x_815); -lean_ctor_set(x_817, 1, x_816); -return x_817; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } -block_852: +case 7: { -lean_object* x_820; -lean_dec(x_819); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_820 = l_Lean_Meta_getFunInfo(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_820) == 0) -{ -lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; uint8_t x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; -x_821 = lean_ctor_get(x_820, 0); -lean_inc(x_821); -x_822 = lean_ctor_get(x_820, 1); -lean_inc(x_822); -lean_dec(x_820); -x_823 = lean_ctor_get(x_821, 0); -lean_inc(x_823); -lean_dec(x_821); -x_824 = lean_array_get_size(x_3); -x_825 = lean_unsigned_to_nat(0u); -x_826 = lean_unsigned_to_nat(1u); -x_827 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_827, 0, x_825); -lean_ctor_set(x_827, 1, x_824); -lean_ctor_set(x_827, 2, x_826); -x_828 = 0; -x_829 = lean_box(x_828); -x_830 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_830, 0, x_3); -lean_ctor_set(x_830, 1, x_829); -lean_inc(x_10); -lean_inc(x_9); +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_1, 1); +lean_inc(x_27); +x_28 = lean_ctor_get(x_1, 2); +lean_inc(x_28); +x_29 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +x_30 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___boxed), 9, 1); +lean_closure_set(x_30, 0, x_1); lean_inc(x_8); lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -lean_inc(x_2); -x_831 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_2, x_823, x_827, x_827, x_830, x_825, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_822); -lean_dec(x_827); -lean_dec(x_823); -if (lean_obj_tag(x_831) == 0) -{ -lean_object* x_832; lean_object* x_833; lean_object* x_834; uint8_t x_835; -x_832 = lean_ctor_get(x_831, 0); -lean_inc(x_832); -x_833 = lean_ctor_get(x_832, 0); -lean_inc(x_833); -x_834 = lean_ctor_get(x_833, 1); -lean_inc(x_834); -x_835 = lean_unbox(x_834); -lean_dec(x_834); -if (x_835 == 0) -{ -lean_object* x_836; lean_object* x_837; lean_object* x_838; -lean_dec(x_833); -lean_dec(x_2); -x_836 = lean_ctor_get(x_831, 1); -lean_inc(x_836); -lean_dec(x_831); -x_837 = lean_ctor_get(x_832, 1); -lean_inc(x_837); -lean_dec(x_832); -lean_inc(x_1); -x_838 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_1, x_5, x_837, x_7, x_8, x_9, x_10, x_836); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_838; -} -else -{ -lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; -x_839 = lean_ctor_get(x_831, 1); -lean_inc(x_839); -lean_dec(x_831); -x_840 = lean_ctor_get(x_832, 1); -lean_inc(x_840); -lean_dec(x_832); -x_841 = lean_ctor_get(x_833, 0); -lean_inc(x_841); -lean_dec(x_833); -x_842 = l_Lean_mkAppN(x_2, x_841); -lean_dec(x_841); -x_843 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_842, x_5, x_840, x_7, x_8, x_9, x_10, x_839); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -return x_843; -} -} -else +lean_inc(x_3); +lean_inc(x_27); +x_31 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_31) == 0) { -uint8_t x_844; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_844 = !lean_is_exclusive(x_831); -if (x_844 == 0) +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_Expr_hasLooseBVars(x_28); +if (x_36 == 0) { -return x_831; -} -else +lean_object* x_37; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_3); +lean_inc(x_28); +x_37 = l_Lean_Meta_Grind_Canon_canonImpl_visit(x_28, x_3, x_35, x_5, x_6, x_7, x_8, x_33); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_845; lean_object* x_846; lean_object* x_847; -x_845 = lean_ctor_get(x_831, 0); -x_846 = lean_ctor_get(x_831, 1); -lean_inc(x_846); -lean_inc(x_845); -lean_dec(x_831); -x_847 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_847, 0, x_845); -lean_ctor_set(x_847, 1, x_846); -return x_847; -} -} +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +lean_dec(x_38); +x_42 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(x_26, x_27, x_28, x_29, x_34, x_30, x_40, x_3, x_41, x_5, x_6, x_7, x_8, x_39); +return x_42; } else { -uint8_t x_848; -lean_dec(x_10); -lean_dec(x_9); +uint8_t x_43; +lean_dec(x_34); +lean_dec(x_30); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_848 = !lean_is_exclusive(x_820); -if (x_848 == 0) +x_43 = !lean_is_exclusive(x_37); +if (x_43 == 0) { -return x_820; +return x_37; } else { -lean_object* x_849; lean_object* x_850; lean_object* x_851; -x_849 = lean_ctor_get(x_820, 0); -x_850 = lean_ctor_get(x_820, 1); -lean_inc(x_850); -lean_inc(x_849); -lean_dec(x_820); -x_851 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_851, 0, x_849); -lean_ctor_set(x_851, 1, x_850); -return x_851; -} -} -} -} +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_37, 0); +x_45 = lean_ctor_get(x_37, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_37); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) +else { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +lean_object* x_47; +lean_inc(x_28); +x_47 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(x_26, x_27, x_28, x_29, x_34, x_30, x_28, x_3, x_35, x_5, x_6, x_7, x_8, x_33); +return x_47; +} } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_ptr_addr(x_4); -x_8 = lean_ptr_addr(x_1); -x_9 = lean_usize_dec_eq(x_7, x_8); -if (x_9 == 0) +uint8_t x_48; +lean_dec(x_30); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_48 = !lean_is_exclusive(x_31); +if (x_48 == 0) { -x_2 = x_6; -goto _start; +return x_31; } else { -lean_object* x_11; -lean_inc(x_5); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_5); -return x_11; -} +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_31, 0); +x_50 = lean_ctor_get(x_31, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_31); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1() { -_start: +default: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +lean_object* x_52; lean_object* x_53; +x_52 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_3); +x_53 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7(x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_10); -x_12 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1; -lean_inc(x_11); -x_13 = lean_mk_array(x_11, x_12); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_11, x_14); -lean_dec(x_11); -lean_inc(x_1); -x_16 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9(x_1, x_1, x_13, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_16; -} +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +lean_dec(x_54); +x_58 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_56, x_3, x_57, x_5, x_6, x_7, x_8, x_55); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_58; } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Canon", 28, 28); -return x_1; -} +uint8_t x_59; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_53); +if (x_59 == 0) +{ +return x_53; } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Canon.canonImpl.visit", 37, 37); -return x_1; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_53, 0); +x_61 = lean_ctor_get(x_53, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_53); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; } } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1; -x_2 = l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2; -x_3 = lean_unsigned_to_nat(115u); -x_4 = lean_unsigned_to_nat(18u); -x_5 = l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_9; lean_object* x_10; -lean_dec(x_1); -x_9 = l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4; -x_10 = l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -case 5: -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_8); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_3, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; size_t x_33; lean_object* x_34; lean_object* x_35; -x_15 = lean_ctor_get(x_11, 1); -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_16); -x_19 = lean_ptr_addr(x_1); -x_20 = lean_usize_to_uint64(x_19); -x_21 = 11; -x_22 = lean_uint64_mix_hash(x_20, x_21); -x_23 = 32; -x_24 = lean_uint64_shift_right(x_22, x_23); -x_25 = lean_uint64_xor(x_22, x_24); -x_26 = 16; -x_27 = lean_uint64_shift_right(x_25, x_26); -x_28 = lean_uint64_xor(x_25, x_27); -x_29 = lean_uint64_to_usize(x_28); -x_30 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_31 = 1; -x_32 = lean_usize_sub(x_30, x_31); -x_33 = lean_usize_land(x_29, x_32); -x_34 = lean_array_uget(x_16, x_33); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; lean_object* x_34; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_12, 1); +x_16 = lean_ctor_get(x_12, 0); lean_dec(x_16); -x_35 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_34); -lean_dec(x_34); -if (lean_obj_tag(x_35) == 0) +x_17 = lean_array_get_size(x_15); +x_18 = lean_ptr_addr(x_1); +x_19 = lean_usize_to_uint64(x_18); +x_20 = 11; +x_21 = lean_uint64_mix_hash(x_19, x_20); +x_22 = 32; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = 16; +x_26 = lean_uint64_shift_right(x_24, x_25); +x_27 = lean_uint64_xor(x_24, x_26); +x_28 = lean_uint64_to_usize(x_27); +x_29 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_30 = 1; +x_31 = lean_usize_sub(x_29, x_30); +x_32 = lean_usize_land(x_28, x_31); +x_33 = lean_array_uget(x_15, x_32); +lean_dec(x_15); +x_34 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_33); +lean_dec(x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_36; lean_object* x_37; -lean_free_object(x_13); -lean_free_object(x_11); -x_36 = lean_box(0); -x_37 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -return x_37; +lean_object* x_35; lean_object* x_36; +lean_free_object(x_12); +lean_free_object(x_10); +x_35 = lean_box(0); +x_36 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(x_1, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_14); +return x_36; } else { -lean_object* x_38; +lean_object* x_37; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_38 = lean_ctor_get(x_35, 0); -lean_inc(x_38); -lean_dec(x_35); -lean_ctor_set(x_13, 1, x_3); -lean_ctor_set(x_13, 0, x_38); -return x_11; +x_37 = lean_ctor_get(x_34, 0); +lean_inc(x_37); +lean_dec(x_34); +lean_ctor_set(x_12, 1, x_4); +lean_ctor_set(x_12, 0, x_37); +return x_10; } } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; uint64_t x_43; uint64_t x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; size_t x_52; size_t x_53; size_t x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; -x_39 = lean_ctor_get(x_11, 1); -x_40 = lean_ctor_get(x_13, 1); -lean_inc(x_40); -lean_dec(x_13); -x_41 = lean_array_get_size(x_40); -x_42 = lean_ptr_addr(x_1); -x_43 = lean_usize_to_uint64(x_42); -x_44 = 11; -x_45 = lean_uint64_mix_hash(x_43, x_44); -x_46 = 32; -x_47 = lean_uint64_shift_right(x_45, x_46); -x_48 = lean_uint64_xor(x_45, x_47); -x_49 = 16; -x_50 = lean_uint64_shift_right(x_48, x_49); -x_51 = lean_uint64_xor(x_48, x_50); -x_52 = lean_uint64_to_usize(x_51); -x_53 = lean_usize_of_nat(x_41); -lean_dec(x_41); -x_54 = 1; -x_55 = lean_usize_sub(x_53, x_54); -x_56 = lean_usize_land(x_52, x_55); -x_57 = lean_array_uget(x_40, x_56); +lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; size_t x_51; size_t x_52; size_t x_53; size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; +x_38 = lean_ctor_get(x_10, 1); +x_39 = lean_ctor_get(x_12, 1); +lean_inc(x_39); +lean_dec(x_12); +x_40 = lean_array_get_size(x_39); +x_41 = lean_ptr_addr(x_1); +x_42 = lean_usize_to_uint64(x_41); +x_43 = 11; +x_44 = lean_uint64_mix_hash(x_42, x_43); +x_45 = 32; +x_46 = lean_uint64_shift_right(x_44, x_45); +x_47 = lean_uint64_xor(x_44, x_46); +x_48 = 16; +x_49 = lean_uint64_shift_right(x_47, x_48); +x_50 = lean_uint64_xor(x_47, x_49); +x_51 = lean_uint64_to_usize(x_50); +x_52 = lean_usize_of_nat(x_40); lean_dec(x_40); -x_58 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_57); -lean_dec(x_57); -if (lean_obj_tag(x_58) == 0) +x_53 = 1; +x_54 = lean_usize_sub(x_52, x_53); +x_55 = lean_usize_land(x_51, x_54); +x_56 = lean_array_uget(x_39, x_55); +lean_dec(x_39); +x_57 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_56); +lean_dec(x_56); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_59; lean_object* x_60; -lean_free_object(x_11); -x_59 = lean_box(0); -x_60 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_39); -return x_60; +lean_object* x_58; lean_object* x_59; +lean_free_object(x_10); +x_58 = lean_box(0); +x_59 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(x_1, x_58, x_3, x_4, x_5, x_6, x_7, x_8, x_38); +return x_59; } else { -lean_object* x_61; lean_object* x_62; +lean_object* x_60; lean_object* x_61; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_61 = lean_ctor_get(x_58, 0); -lean_inc(x_61); -lean_dec(x_58); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_3); -lean_ctor_set(x_11, 0, x_62); -return x_11; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +lean_dec(x_57); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_4); +lean_ctor_set(x_10, 0, x_61); +return x_10; } } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; size_t x_68; uint64_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; uint64_t x_76; uint64_t x_77; size_t x_78; size_t x_79; size_t x_80; size_t x_81; size_t x_82; lean_object* x_83; lean_object* x_84; -x_63 = lean_ctor_get(x_11, 0); -x_64 = lean_ctor_get(x_11, 1); -lean_inc(x_64); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; size_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; uint64_t x_76; size_t x_77; size_t x_78; size_t x_79; size_t x_80; size_t x_81; lean_object* x_82; lean_object* x_83; +x_62 = lean_ctor_get(x_10, 0); +x_63 = lean_ctor_get(x_10, 1); lean_inc(x_63); -lean_dec(x_11); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_66 = x_63; -} else { - lean_dec_ref(x_63); - x_66 = lean_box(0); -} -x_67 = lean_array_get_size(x_65); -x_68 = lean_ptr_addr(x_1); -x_69 = lean_usize_to_uint64(x_68); -x_70 = 11; -x_71 = lean_uint64_mix_hash(x_69, x_70); -x_72 = 32; -x_73 = lean_uint64_shift_right(x_71, x_72); -x_74 = lean_uint64_xor(x_71, x_73); -x_75 = 16; -x_76 = lean_uint64_shift_right(x_74, x_75); -x_77 = lean_uint64_xor(x_74, x_76); -x_78 = lean_uint64_to_usize(x_77); -x_79 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_80 = 1; -x_81 = lean_usize_sub(x_79, x_80); -x_82 = lean_usize_land(x_78, x_81); -x_83 = lean_array_uget(x_65, x_82); -lean_dec(x_65); -x_84 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_83); -lean_dec(x_83); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_85; lean_object* x_86; +lean_inc(x_62); +lean_dec(x_10); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_65 = x_62; +} else { + lean_dec_ref(x_62); + x_65 = lean_box(0); +} +x_66 = lean_array_get_size(x_64); +x_67 = lean_ptr_addr(x_1); +x_68 = lean_usize_to_uint64(x_67); +x_69 = 11; +x_70 = lean_uint64_mix_hash(x_68, x_69); +x_71 = 32; +x_72 = lean_uint64_shift_right(x_70, x_71); +x_73 = lean_uint64_xor(x_70, x_72); +x_74 = 16; +x_75 = lean_uint64_shift_right(x_73, x_74); +x_76 = lean_uint64_xor(x_73, x_75); +x_77 = lean_uint64_to_usize(x_76); +x_78 = lean_usize_of_nat(x_66); lean_dec(x_66); -x_85 = lean_box(0); -x_86 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_64); -return x_86; +x_79 = 1; +x_80 = lean_usize_sub(x_78, x_79); +x_81 = lean_usize_land(x_77, x_80); +x_82 = lean_array_uget(x_64, x_81); +lean_dec(x_64); +x_83 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_82); +lean_dec(x_82); +if (lean_obj_tag(x_83) == 0) +{ +lean_object* x_84; lean_object* x_85; +lean_dec(x_65); +x_84 = lean_box(0); +x_85 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(x_1, x_84, x_3, x_4, x_5, x_6, x_7, x_8, x_63); +return x_85; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_87 = lean_ctor_get(x_84, 0); -lean_inc(x_87); -lean_dec(x_84); -if (lean_is_scalar(x_66)) { - x_88 = lean_alloc_ctor(0, 2, 0); +x_86 = lean_ctor_get(x_83, 0); +lean_inc(x_86); +lean_dec(x_83); +if (lean_is_scalar(x_65)) { + x_87 = lean_alloc_ctor(0, 2, 0); } else { - x_88 = x_66; + x_87 = x_65; } +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_4); +x_88 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_3); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_64); -return x_89; +lean_ctor_set(x_88, 1, x_63); +return x_88; } } } -default: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_90; lean_object* x_91; +uint8_t x_9; +x_9 = l_Lean_Expr_isApp(x_1); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = l_Lean_Expr_isForall(x_1); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_1); -lean_ctor_set(x_90, 1, x_3); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_8); -return x_91; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_3); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(0); +x_14 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(x_1, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_14; } } +else +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(x_1, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_16; +} } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__2(x_1, x_2); +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -11036,22 +14373,51 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__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_16; -x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_object* x_18; +x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_16; +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -11060,21 +14426,30 @@ lean_dec(x_3); return x_10; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__10(x_1, x_2); +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__2(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_2); -lean_dec(x_1); -return x_3; +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_2); return x_10; } @@ -11279,12 +14654,12 @@ l_Lean_Meta_Grind_Canon_canonInst___closed__1 = _init_l_Lean_Meta_Grind_Canon_ca l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Canon_0__Lean_Meta_Grind_Canon_ShouldCanonResult_noConfusion___rarg___closed__1); l_Lean_Meta_Grind_Canon_instInhabitedShouldCanonResult = _init_l_Lean_Meta_Grind_Canon_instInhabitedShouldCanonResult(); -l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__1); -l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__2); -l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__1___closed__3); +l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__1); +l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__2); +l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3 = _init_l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__7___closed__3); l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1(); lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__1); l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__2(); @@ -11293,16 +14668,16 @@ l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___cl lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__3); l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4(); lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_Canon_canonImpl_visit___spec__9___closed__4); -l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__1___closed__1); -l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__1); -l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__2); -l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__3); -l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___closed__4); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__1); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__2); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__3); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__4); +l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl_visit___lambda__3___closed__5); l_Lean_Meta_Grind_Canon_canonImpl___closed__1 = _init_l_Lean_Meta_Grind_Canon_canonImpl___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_Canon_canonImpl___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c index 85a332d356ac..381b145afbf4 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Cases.c @@ -13,65 +13,113 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8___boxed(lean_object**); lean_object* l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6___boxed(lean_object**); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_casesOnSuffix; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_RecursorInfo_numIndices(lean_object*); +static lean_object* l_Lean_Meta_Grind_cases___lambda__13___closed__4; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13___boxed(lean_object**); lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7___boxed(lean_object**); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l_Lean_Meta_Grind_cases___lambda__13___closed__3; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_generalizeIndices_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); -lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_cases___lambda__1___closed__1; +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -lean_object* l_Lean_MVarId_clear(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); -lean_object* l_Lean_Meta_getMajorTypeIndices(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_cases___lambda__12___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkRecursorAppPrefix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_cases___lambda__13___closed__1; lean_object* l_Lean_Meta_mkRecursorInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9___boxed(lean_object**); lean_object* l_Lean_Meta_RecursorInfo_numMinors(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__6; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___boxed(lean_object**); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* lean_array_mk(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_tryClearMany(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_cases___lambda__13___closed__2; lean_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2___boxed(lean_object**); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); -lean_object* l_Lean_Meta_generalizeIndices(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12___boxed(lean_object**); static lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__7; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__3; lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__1() { _start: { @@ -135,29 +183,28 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_10 = l_Lean_Expr_fvar___override(x_2); -x_11 = l_Lean_MessageData_ofExpr(x_10); -x_12 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__5; -x_13 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -x_14 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__7; -x_15 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -x_16 = l_Lean_indentExpr(x_4); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_10 = l_Lean_MessageData_ofExpr(x_2); +x_11 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__5; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__7; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Lean_indentExpr(x_4); +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); x_17 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_13); +x_18 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_14); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; -x_21 = l_Lean_Meta_throwTacticEx___rarg(x_20, x_1, x_19, x_5, x_6, x_7, x_8, x_9); -return x_21; +x_19 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_20 = l_Lean_Meta_throwTacticEx___rarg(x_19, x_1, x_18, x_5, x_6, x_7, x_8, x_9); +return x_20; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases_throwInductiveExpected___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -172,25 +219,6 @@ lean_dec(x_5); return x_10; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_array_push(x_3, x_4); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_2); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_9); -return x_14; -} -} static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__1() { _start: { @@ -228,986 +256,10850 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_6, 1); -x_17 = lean_nat_dec_lt(x_8, x_16); -if (x_17 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) { -lean_object* x_18; +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); -lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_15); -return x_18; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_ctor_get(x_7, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_7, 1); -lean_inc(x_27); -lean_dec(x_7); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -x_30 = lean_whnf(x_29, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -if (lean_obj_tag(x_31) == 7) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_ctor_get(x_31, 1); +x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); -x_34 = lean_ctor_get(x_31, 2); +x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); -lean_dec(x_31); -lean_inc(x_11); -lean_inc(x_3); -x_35 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_33, x_3, x_11, x_12, x_13, x_14, x_32); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); lean_dec(x_35); -lean_inc(x_36); -x_38 = l_Lean_Expr_app___override(x_28, x_36); -x_39 = l_Lean_Expr_mvarId_x21(x_36); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_2); -x_40 = l_Lean_MVarId_clear(x_39, x_2, x_11, x_12, x_13, x_14, x_37); -if (lean_obj_tag(x_40) == 0) +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1(x_38, x_34, x_26, x_41, x_11, x_12, x_13, x_14, x_42); -x_44 = lean_ctor_get(x_43, 0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_19 = x_44; -x_20 = x_45; -goto block_25; -} -else -{ -uint8_t x_46; -lean_dec(x_38); -lean_dec(x_34); -lean_dec(x_26); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_46 = !lean_is_exclusive(x_40); -if (x_46 == 0) -{ -return x_40; -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_40, 0); -x_48 = lean_ctor_get(x_40, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_40); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; -} -} -} -else +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -lean_dec(x_31); -lean_dec(x_28); -lean_dec(x_26); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -x_50 = lean_ctor_get(x_30, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); -lean_dec(x_30); -x_51 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; -x_52 = l_Lean_Meta_throwTacticEx___rarg(x_4, x_1, x_51, x_11, x_12, x_13, x_14, x_50); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) -{ -return x_52; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_52, 0); -x_55 = lean_ctor_get(x_52, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_52); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; -} -} +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; } else { uint8_t x_57; -lean_dec(x_28); -lean_dec(x_26); +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); -lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_57 = !lean_is_exclusive(x_30); +x_57 = !lean_is_exclusive(x_49); if (x_57 == 0) { -return x_30; +return x_49; } else { lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_30, 0); -x_59 = lean_ctor_get(x_30, 1); +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); lean_inc(x_59); lean_inc(x_58); -lean_dec(x_30); +lean_dec(x_49); x_60 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_60, 0, x_58); lean_ctor_set(x_60, 1, x_59); return x_60; } } -block_25: -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_6, 2); -x_23 = lean_nat_add(x_8, x_22); -lean_dec(x_8); -x_7 = x_21; -x_8 = x_23; -x_9 = lean_box(0); -x_10 = lean_box(0); -x_15 = x_20; -goto _start; -} -} } -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +else { -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_5, 1); -x_16 = lean_nat_dec_lt(x_7, x_15); -if (x_16 == 0) +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_17; -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_14); -return x_17; +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} } else { -lean_object* x_18; lean_object* x_19; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_6, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_6, 1); -lean_inc(x_26); -lean_dec(x_6); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_29 = lean_whnf(x_28, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_29) == 0) +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) { -lean_object* x_30; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -if (lean_obj_tag(x_30) == 7) +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); lean_dec(x_29); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -x_33 = lean_ctor_get(x_30, 2); -lean_inc(x_33); -lean_dec(x_30); -lean_inc(x_10); -lean_inc(x_1); -x_34 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_32, x_1, x_10, x_11, x_12, x_13, x_31); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -lean_inc(x_35); -x_37 = l_Lean_Expr_app___override(x_27, x_35); -x_38 = l_Lean_Expr_mvarId_x21(x_35); -lean_dec(x_35); -x_39 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1(x_37, x_33, x_25, x_38, x_10, x_11, x_12, x_13, x_36); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_18 = x_40; -x_19 = x_41; -goto block_24; +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -lean_dec(x_30); -lean_dec(x_27); -lean_dec(x_25); -lean_dec(x_7); -lean_dec(x_1); -x_42 = lean_ctor_get(x_29, 1); -lean_inc(x_42); +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); lean_dec(x_29); -x_43 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; -x_44 = l_Lean_Meta_throwTacticEx___rarg(x_3, x_2, x_43, x_10, x_11, x_12, x_13, x_42); -lean_dec(x_13); -lean_dec(x_12); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_11); -lean_dec(x_10); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) { -return x_44; +return x_87; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_44, 0); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_44); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; } } } else { -uint8_t x_49; -lean_dec(x_27); -lean_dec(x_25); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_49 = !lean_is_exclusive(x_29); -if (x_49 == 0) +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) { -return x_29; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_29, 0); -x_51 = lean_ctor_get(x_29, 1); -lean_inc(x_51); -lean_inc(x_50); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_2); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_2, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_4); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_4); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_2); +x_79 = l_Lean_Name_num___override(x_2, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_4); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_4); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_1, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_3); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_3); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_3); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_3); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_3); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_4, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_nat_dec_lt(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_35 = lean_whnf(x_34, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 7) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 2); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_dec_lt(x_40, x_7); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +lean_inc(x_14); +lean_inc(x_1); +x_42 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_1, x_14, x_15, x_16, x_17, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +x_46 = l_Lean_Expr_app___override(x_33, x_44); +x_47 = l_Lean_Expr_mvarId_x21(x_44); +lean_dec(x_44); +lean_inc(x_4); +lean_inc(x_5); +x_48 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_49 = l_Lean_MVarId_tryClearMany(x_47, x_48, x_14, x_15, x_16, x_17, x_45); +lean_dec(x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = lean_array_push(x_31, x_50); +x_53 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 0, x_46); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_42); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_22 = x_56; +x_23 = x_51; +goto block_28; +} +else +{ +uint8_t x_57; +lean_dec(x_46); +lean_free_object(x_42); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_49); +if (x_57 == 0) +{ +return x_49; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_49, 0); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_42, 0); +x_62 = lean_ctor_get(x_42, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_42); +lean_inc(x_61); +x_63 = l_Lean_Expr_app___override(x_33, x_61); +x_64 = l_Lean_Expr_mvarId_x21(x_61); +lean_dec(x_61); +lean_inc(x_4); +lean_inc(x_5); +x_65 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_66 = l_Lean_MVarId_tryClearMany(x_64, x_65, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_array_push(x_31, x_67); +x_70 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_63); +lean_ctor_set(x_71, 1, x_39); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_22 = x_74; +x_23 = x_68; +goto block_28; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_63); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_66, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_77 = x_66; +} else { + lean_dec_ref(x_66); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_inc(x_29); +lean_inc(x_1); +x_79 = l_Lean_Name_num___override(x_1, x_29); +lean_inc(x_14); +x_80 = l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(x_38, x_79, x_14, x_15, x_16, x_17, x_37); +x_81 = !lean_is_exclusive(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_80, 0); +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_84 = l_Lean_Expr_app___override(x_33, x_82); +x_85 = l_Lean_Expr_mvarId_x21(x_82); +lean_dec(x_82); +lean_inc(x_4); +lean_inc(x_5); +x_86 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_87 = l_Lean_MVarId_tryClearMany(x_85, x_86, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_array_push(x_31, x_88); +x_91 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +lean_ctor_set(x_80, 1, x_39); +lean_ctor_set(x_80, 0, x_84); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_80); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_22 = x_94; +x_23 = x_89; +goto block_28; +} +else +{ +uint8_t x_95; +lean_dec(x_84); +lean_free_object(x_80); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_87); +if (x_95 == 0) +{ +return x_87; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_87, 0); +x_97 = lean_ctor_get(x_87, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_87); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_80, 0); +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_80); +lean_inc(x_99); +x_101 = l_Lean_Expr_app___override(x_33, x_99); +x_102 = l_Lean_Expr_mvarId_x21(x_99); +lean_dec(x_99); +lean_inc(x_4); +lean_inc(x_5); +x_103 = lean_array_push(x_5, x_4); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_104 = l_Lean_MVarId_tryClearMany(x_102, x_103, x_14, x_15, x_16, x_17, x_100); +lean_dec(x_103); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = lean_array_push(x_31, x_105); +x_108 = lean_nat_add(x_29, x_40); +lean_dec(x_29); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_101); +lean_ctor_set(x_109, 1, x_39); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_22 = x_112; +x_23 = x_106; +goto block_28; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_101); +lean_dec(x_39); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_113 = lean_ctor_get(x_104, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_104, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_115 = x_104; +} else { + lean_dec_ref(x_104); + x_115 = lean_box(0); +} +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); +} else { + x_116 = x_115; +} +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_36); +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_117 = lean_ctor_get(x_35, 1); +lean_inc(x_117); +lean_dec(x_35); +x_118 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4; +x_119 = l_Lean_Meta_throwTacticEx___rarg(x_6, x_3, x_118, x_14, x_15, x_16, x_17, x_117); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +return x_119; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_29); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_35); +if (x_124 == 0) +{ +return x_35; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_35, 0); +x_126 = lean_ctor_get(x_35, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_35); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +block_28: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_nat_add(x_11, x_25); +lean_dec(x_11); +x_10 = x_24; +x_11 = x_26; +x_12 = lean_box(0); +x_13 = lean_box(0); +x_18 = x_23; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +lean_inc(x_6); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_27); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_29 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12(x_7, x_4, x_3, x_1, x_6, x_2, x_22, x_25, x_25, x_28, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_35, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = lean_array_to_list(x_34); +lean_ctor_set(x_36, 0, x_39); +return x_36; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +lean_dec(x_36); +x_41 = lean_array_to_list(x_34); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_29); +if (x_43 == 0) +{ +return x_29; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_29, 0); +x_45 = lean_ctor_get(x_29, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_29); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_19); +if (x_47 == 0) +{ +return x_19; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_19); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_13); +if (x_51 == 0) +{ +return x_13; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_13, 0); +x_53 = lean_ctor_get(x_13, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_13); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__12___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_mkAppN(x_14, x_5); +lean_inc(x_3); +x_17 = l_Lean_Expr_fvar___override(x_3); +x_18 = l_Lean_Expr_app___override(x_16, x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_18); +x_19 = lean_infer_type(x_18, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Meta_RecursorInfo_numMinors(x_4); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_20); +x_27 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_24); +lean_ctor_set(x_29, 1, x_28); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_30 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13(x_6, x_4, x_1, x_3, x_7, x_2, x_22, x_25, x_25, x_29, x_23, lean_box(0), lean_box(0), x_8, x_9, x_10, x_11, x_21); +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_4); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +x_34 = lean_ctor_get(x_30, 1); +lean_inc(x_34); +lean_dec(x_30); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_36 = lean_ctor_get(x_33, 0); +lean_inc(x_36); +lean_dec(x_33); +x_37 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_36, x_8, x_9, x_10, x_11, x_34); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +x_40 = lean_array_to_list(x_35); +lean_ctor_set(x_37, 0, x_40); +return x_37; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 1); +lean_inc(x_41); +lean_dec(x_37); +x_42 = lean_array_to_list(x_35); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +return x_43; +} +} +else +{ +uint8_t x_44; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_44 = !lean_is_exclusive(x_30); +if (x_44 == 0) +{ +return x_30; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_30, 0); +x_46 = lean_ctor_get(x_30, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_30); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +uint8_t x_48; +lean_dec(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_19); +if (x_48 == 0) +{ +return x_19; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_19, 0); +x_50 = lean_ctor_get(x_19, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_19); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +} +else +{ +uint8_t x_52; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_52 = !lean_is_exclusive(x_13); +if (x_52 == 0) +{ +return x_13; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_13, 0); +x_54 = lean_ctor_get(x_13, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_13); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__13___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__13___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_cases___lambda__13___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static size_t _init_l_Lean_Meta_Grind_cases___lambda__13___closed__3() { +_start: +{ +lean_object* x_1; size_t x_2; +x_1 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +x_2 = lean_array_size(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__13___closed__4() { +_start: +{ +size_t x_1; size_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Lean_Meta_Grind_cases___lambda__13___closed__3; +x_3 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +x_4 = l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(x_2, x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +lean_inc(x_1); +x_8 = l_Lean_MVarId_getTag(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_11 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_14 = lean_whnf(x_12, x_3, x_4, x_5, x_6, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Expr_getAppFn(x_15); +if (lean_obj_tag(x_17) == 4) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +lean_inc(x_18); +x_19 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_18, x_3, x_4, x_5, x_6, x_16); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 5) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_casesOnSuffix; +x_23 = l_Lean_Name_str___override(x_18, x_22); +x_24 = lean_box(0); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_25 = l_Lean_Meta_mkRecursorInfo(x_23, x_24, x_3, x_4, x_5, x_6, x_21); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_Meta_RecursorInfo_numIndices(x_26); +x_29 = lean_unsigned_to_nat(0u); +x_30 = lean_nat_dec_lt(x_29, x_28); +lean_dec(x_28); +if (x_30 == 0) +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_32 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_31, x_5, x_6, x_27); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_35 = l_Lean_MVarId_assert(x_1, x_33, x_15, x_2, x_3, x_4, x_5, x_6, x_34); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_39 = l_Lean_Meta_intro1Core(x_36, x_38, x_3, x_4, x_5, x_6, x_37); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_45 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_46 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_43); +x_47 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__1___boxed), 12, 7); +lean_closure_set(x_47, 0, x_43); +lean_closure_set(x_47, 1, x_44); +lean_closure_set(x_47, 2, x_42); +lean_closure_set(x_47, 3, x_26); +lean_closure_set(x_47, 4, x_45); +lean_closure_set(x_47, 5, x_46); +lean_closure_set(x_47, 6, x_9); +x_48 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_43, x_47, x_3, x_4, x_5, x_6, x_41); +return x_48; +} +else +{ +uint8_t x_49; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_49 = !lean_is_exclusive(x_39); +if (x_49 == 0) +{ +return x_39; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_39, 0); +x_51 = lean_ctor_get(x_39, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_39); x_52 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); return x_52; } } -block_24: +} +else +{ +uint8_t x_53; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_53 = !lean_is_exclusive(x_35); +if (x_53 == 0) +{ +return x_35; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_35, 0); +x_55 = lean_ctor_get(x_35, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_35); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; +} +} +} +case 1: +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_15); +x_57 = lean_ctor_get(x_2, 0); +lean_inc(x_57); +lean_dec(x_2); +x_58 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_59 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_26); +lean_inc(x_57); +lean_inc(x_1); +x_60 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_58, x_57, x_26, x_59, x_3, x_4, x_5, x_6, x_27); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_mkAppN(x_61, x_59); +lean_inc(x_57); +x_64 = l_Lean_Expr_fvar___override(x_57); +x_65 = l_Lean_Expr_app___override(x_63, x_64); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_65); +x_66 = lean_infer_type(x_65, x_3, x_4, x_5, x_6, x_62); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = l_Lean_Meta_RecursorInfo_numMinors(x_26); +x_70 = lean_unsigned_to_nat(1u); +lean_inc(x_69); +x_71 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_71, 0, x_29); +lean_ctor_set(x_71, 1, x_69); +lean_ctor_set(x_71, 2, x_70); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_65); +lean_ctor_set(x_72, 1, x_67); +x_73 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_72); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_70); +lean_ctor_set(x_75, 1, x_74); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +x_76 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(x_1, x_9, x_26, x_57, x_73, x_58, x_69, x_71, x_71, x_75, x_29, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_68); +lean_dec(x_71); +lean_dec(x_69); +lean_dec(x_26); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +x_80 = lean_ctor_get(x_76, 1); +lean_inc(x_80); +lean_dec(x_76); +x_81 = lean_ctor_get(x_78, 0); +lean_inc(x_81); +lean_dec(x_78); +x_82 = lean_ctor_get(x_79, 0); +lean_inc(x_82); +lean_dec(x_79); +x_83 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_82, x_3, x_4, x_5, x_6, x_80); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_84 = !lean_is_exclusive(x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; +x_85 = lean_ctor_get(x_83, 0); +lean_dec(x_85); +x_86 = lean_array_to_list(x_81); +lean_ctor_set(x_83, 0, x_86); +return x_83; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_83, 1); +lean_inc(x_87); +lean_dec(x_83); +x_88 = lean_array_to_list(x_81); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_87); +return x_89; +} +} +else +{ +uint8_t x_90; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_90 = !lean_is_exclusive(x_76); +if (x_90 == 0) +{ +return x_76; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_76, 0); +x_92 = lean_ctor_get(x_76, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_76); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; +} +} +} +else +{ +uint8_t x_94; +lean_dec(x_65); +lean_dec(x_57); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_94 = !lean_is_exclusive(x_66); +if (x_94 == 0) +{ +return x_66; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_66, 0); +x_96 = lean_ctor_get(x_66, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_66); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; +} +} +} +else +{ +uint8_t x_98; +lean_dec(x_57); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_98 = !lean_is_exclusive(x_60); +if (x_98 == 0) +{ +return x_60; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_60, 0); +x_100 = lean_ctor_get(x_60, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_60); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; +} +} +} +case 2: +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_102 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_103 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_102, x_5, x_6, x_27); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_106 = l_Lean_MVarId_assert(x_1, x_104, x_15, x_2, x_3, x_4, x_5, x_6, x_105); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_110 = l_Lean_Meta_intro1Core(x_107, x_109, x_3, x_4, x_5, x_6, x_108); +if (lean_obj_tag(x_110) == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_110, 1); +lean_inc(x_112); +lean_dec(x_110); +x_113 = lean_ctor_get(x_111, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_111, 1); +lean_inc(x_114); +lean_dec(x_111); +x_115 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_116 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_117 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_114); +x_118 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__2___boxed), 12, 7); +lean_closure_set(x_118, 0, x_114); +lean_closure_set(x_118, 1, x_115); +lean_closure_set(x_118, 2, x_113); +lean_closure_set(x_118, 3, x_26); +lean_closure_set(x_118, 4, x_116); +lean_closure_set(x_118, 5, x_117); +lean_closure_set(x_118, 6, x_9); +x_119 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_114, x_118, x_3, x_4, x_5, x_6, x_112); +return x_119; +} +else +{ +uint8_t x_120; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_120 = !lean_is_exclusive(x_110); +if (x_120 == 0) +{ +return x_110; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_110, 0); +x_122 = lean_ctor_get(x_110, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_110); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_124 = !lean_is_exclusive(x_106); +if (x_124 == 0) +{ +return x_106; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_106, 0); +x_126 = lean_ctor_get(x_106, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_106); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +} +case 3: +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_128 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_129 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_128, x_5, x_6, x_27); +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_129, 1); +lean_inc(x_131); +lean_dec(x_129); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_132 = l_Lean_MVarId_assert(x_1, x_130, x_15, x_2, x_3, x_4, x_5, x_6, x_131); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; lean_object* x_134; uint8_t x_135; lean_object* x_136; +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_132, 1); +lean_inc(x_134); +lean_dec(x_132); +x_135 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_136 = l_Lean_Meta_intro1Core(x_133, x_135, x_3, x_4, x_5, x_6, x_134); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +lean_dec(x_136); +x_139 = lean_ctor_get(x_137, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_137, 1); +lean_inc(x_140); +lean_dec(x_137); +x_141 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_142 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_143 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_140); +x_144 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__3___boxed), 12, 7); +lean_closure_set(x_144, 0, x_140); +lean_closure_set(x_144, 1, x_141); +lean_closure_set(x_144, 2, x_139); +lean_closure_set(x_144, 3, x_26); +lean_closure_set(x_144, 4, x_142); +lean_closure_set(x_144, 5, x_143); +lean_closure_set(x_144, 6, x_9); +x_145 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_140, x_144, x_3, x_4, x_5, x_6, x_138); +return x_145; +} +else +{ +uint8_t x_146; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_146 = !lean_is_exclusive(x_136); +if (x_146 == 0) +{ +return x_136; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_136, 0); +x_148 = lean_ctor_get(x_136, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_136); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +return x_149; +} +} +} +else +{ +uint8_t x_150; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_150 = !lean_is_exclusive(x_132); +if (x_150 == 0) +{ +return x_132; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_132, 0); +x_152 = lean_ctor_get(x_132, 1); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_132); +x_153 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_153, 0, x_151); +lean_ctor_set(x_153, 1, x_152); +return x_153; +} +} +} +case 4: +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_154 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_155 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_154, x_5, x_6, x_27); +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec(x_155); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_158 = l_Lean_MVarId_assert(x_1, x_156, x_15, x_2, x_3, x_4, x_5, x_6, x_157); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +lean_dec(x_158); +x_161 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_162 = l_Lean_Meta_intro1Core(x_159, x_161, x_3, x_4, x_5, x_6, x_160); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = lean_ctor_get(x_163, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_163, 1); +lean_inc(x_166); +lean_dec(x_163); +x_167 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_168 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_169 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_166); +x_170 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__4___boxed), 12, 7); +lean_closure_set(x_170, 0, x_166); +lean_closure_set(x_170, 1, x_167); +lean_closure_set(x_170, 2, x_165); +lean_closure_set(x_170, 3, x_26); +lean_closure_set(x_170, 4, x_168); +lean_closure_set(x_170, 5, x_169); +lean_closure_set(x_170, 6, x_9); +x_171 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_166, x_170, x_3, x_4, x_5, x_6, x_164); +return x_171; +} +else +{ +uint8_t x_172; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_172 = !lean_is_exclusive(x_162); +if (x_172 == 0) +{ +return x_162; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_173 = lean_ctor_get(x_162, 0); +x_174 = lean_ctor_get(x_162, 1); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_162); +x_175 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_175, 0, x_173); +lean_ctor_set(x_175, 1, x_174); +return x_175; +} +} +} +else +{ +uint8_t x_176; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_176 = !lean_is_exclusive(x_158); +if (x_176 == 0) +{ +return x_158; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_177 = lean_ctor_get(x_158, 0); +x_178 = lean_ctor_get(x_158, 1); +lean_inc(x_178); +lean_inc(x_177); +lean_dec(x_158); +x_179 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_179, 0, x_177); +lean_ctor_set(x_179, 1, x_178); +return x_179; +} +} +} +case 5: +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_180 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_181 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_180, x_5, x_6, x_27); +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_184 = l_Lean_MVarId_assert(x_1, x_182, x_15, x_2, x_3, x_4, x_5, x_6, x_183); +if (lean_obj_tag(x_184) == 0) +{ +lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_184, 1); +lean_inc(x_186); +lean_dec(x_184); +x_187 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_188 = l_Lean_Meta_intro1Core(x_185, x_187, x_3, x_4, x_5, x_6, x_186); +if (lean_obj_tag(x_188) == 0) +{ +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +lean_dec(x_188); +x_191 = lean_ctor_get(x_189, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_189, 1); +lean_inc(x_192); +lean_dec(x_189); +x_193 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_194 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_195 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_192); +x_196 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__5___boxed), 12, 7); +lean_closure_set(x_196, 0, x_192); +lean_closure_set(x_196, 1, x_193); +lean_closure_set(x_196, 2, x_191); +lean_closure_set(x_196, 3, x_26); +lean_closure_set(x_196, 4, x_194); +lean_closure_set(x_196, 5, x_195); +lean_closure_set(x_196, 6, x_9); +x_197 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_192, x_196, x_3, x_4, x_5, x_6, x_190); +return x_197; +} +else +{ +uint8_t x_198; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_198 = !lean_is_exclusive(x_188); +if (x_198 == 0) +{ +return x_188; +} +else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_5, 2); -x_22 = lean_nat_add(x_7, x_21); -lean_dec(x_7); -x_6 = x_20; -x_7 = x_22; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_19; -goto _start; +lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_199 = lean_ctor_get(x_188, 0); +x_200 = lean_ctor_get(x_188, 1); +lean_inc(x_200); +lean_inc(x_199); +lean_dec(x_188); +x_201 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_201, 0, x_199); +lean_ctor_set(x_201, 1, x_200); +return x_201; } } } +else +{ +uint8_t x_202; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_202 = !lean_is_exclusive(x_184); +if (x_202 == 0) +{ +return x_184; } -static lean_object* _init_l_Lean_Meta_Grind_cases___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; +lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_203 = lean_ctor_get(x_184, 0); +x_204 = lean_ctor_get(x_184, 1); +lean_inc(x_204); +lean_inc(x_203); +lean_dec(x_184); +x_205 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +return x_205; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +} +case 6: { -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_206 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_207 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_206, x_5, x_6, x_27); +x_208 = lean_ctor_get(x_207, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_207, 1); +lean_inc(x_209); +lean_dec(x_207); +lean_inc(x_6); +lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_12 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_2, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) +x_210 = l_Lean_MVarId_assert(x_1, x_208, x_15, x_2, x_3, x_4, x_5, x_6, x_209); +if (lean_obj_tag(x_210) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_mkAppN(x_13, x_5); -x_16 = l_Lean_Expr_fvar___override(x_3); -x_17 = l_Lean_Expr_app___override(x_15, x_16); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_17); -x_18 = lean_infer_type(x_17, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_18) == 0) +lean_object* x_211; lean_object* x_212; uint8_t x_213; lean_object* x_214; +x_211 = lean_ctor_get(x_210, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_210, 1); +lean_inc(x_212); +lean_dec(x_210); +x_213 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_214 = l_Lean_Meta_intro1Core(x_211, x_213, x_3, x_4, x_5, x_6, x_212); +if (lean_obj_tag(x_214) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Meta_RecursorInfo_numMinors(x_4); -lean_dec(x_4); -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_17); -lean_ctor_set(x_25, 1, x_19); -x_26 = l_Lean_Meta_Grind_cases___lambda__1___closed__1; -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_1); -x_28 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(x_6, x_1, x_2, x_24, x_24, x_27, x_22, lean_box(0), lean_box(0), x_7, x_8, x_9, x_10, x_20); -lean_dec(x_24); -if (lean_obj_tag(x_28) == 0) +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_214, 1); +lean_inc(x_216); +lean_dec(x_214); +x_217 = lean_ctor_get(x_215, 0); +lean_inc(x_217); +x_218 = lean_ctor_get(x_215, 1); +lean_inc(x_218); +lean_dec(x_215); +x_219 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_220 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_221 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_218); +x_222 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__6___boxed), 12, 7); +lean_closure_set(x_222, 0, x_218); +lean_closure_set(x_222, 1, x_219); +lean_closure_set(x_222, 2, x_217); +lean_closure_set(x_222, 3, x_26); +lean_closure_set(x_222, 4, x_220); +lean_closure_set(x_222, 5, x_221); +lean_closure_set(x_222, 6, x_9); +x_223 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_218, x_222, x_3, x_4, x_5, x_6, x_216); +return x_223; +} +else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_ctor_get(x_29, 0); -lean_inc(x_32); -lean_dec(x_29); -x_33 = lean_ctor_get(x_30, 0); -lean_inc(x_33); -lean_dec(x_30); -x_34 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_33, x_7, x_8, x_9, x_10, x_31); -lean_dec(x_10); +uint8_t x_224; +lean_dec(x_26); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_224 = !lean_is_exclusive(x_214); +if (x_224 == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -x_37 = lean_array_to_list(x_32); -lean_ctor_set(x_34, 0, x_37); -return x_34; +return x_214; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_34, 1); -lean_inc(x_38); -lean_dec(x_34); -x_39 = lean_array_to_list(x_32); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -return x_40; +lean_object* x_225; lean_object* x_226; lean_object* x_227; +x_225 = lean_ctor_get(x_214, 0); +x_226 = lean_ctor_get(x_214, 1); +lean_inc(x_226); +lean_inc(x_225); +lean_dec(x_214); +x_227 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_227, 0, x_225); +lean_ctor_set(x_227, 1, x_226); +return x_227; +} } } else { -uint8_t x_41; -lean_dec(x_10); +uint8_t x_228; +lean_dec(x_26); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_1); -x_41 = !lean_is_exclusive(x_28); -if (x_41 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_228 = !lean_is_exclusive(x_210); +if (x_228 == 0) { -return x_28; +return x_210; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_28, 0); -x_43 = lean_ctor_get(x_28, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_28); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_229 = lean_ctor_get(x_210, 0); +x_230 = lean_ctor_get(x_210, 1); +lean_inc(x_230); +lean_inc(x_229); +lean_dec(x_210); +x_231 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_231, 0, x_229); +lean_ctor_set(x_231, 1, x_230); +return x_231; +} } } +case 7: +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; +x_232 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_233 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_232, x_5, x_6, x_27); +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_233, 1); +lean_inc(x_235); +lean_dec(x_233); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_236 = l_Lean_MVarId_assert(x_1, x_234, x_15, x_2, x_3, x_4, x_5, x_6, x_235); +if (lean_obj_tag(x_236) == 0) +{ +lean_object* x_237; lean_object* x_238; uint8_t x_239; lean_object* x_240; +x_237 = lean_ctor_get(x_236, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_236, 1); +lean_inc(x_238); +lean_dec(x_236); +x_239 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_240 = l_Lean_Meta_intro1Core(x_237, x_239, x_3, x_4, x_5, x_6, x_238); +if (lean_obj_tag(x_240) == 0) +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_241 = lean_ctor_get(x_240, 0); +lean_inc(x_241); +x_242 = lean_ctor_get(x_240, 1); +lean_inc(x_242); +lean_dec(x_240); +x_243 = lean_ctor_get(x_241, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_241, 1); +lean_inc(x_244); +lean_dec(x_241); +x_245 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_246 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_247 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_244); +x_248 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__7___boxed), 12, 7); +lean_closure_set(x_248, 0, x_244); +lean_closure_set(x_248, 1, x_245); +lean_closure_set(x_248, 2, x_243); +lean_closure_set(x_248, 3, x_26); +lean_closure_set(x_248, 4, x_246); +lean_closure_set(x_248, 5, x_247); +lean_closure_set(x_248, 6, x_9); +x_249 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_244, x_248, x_3, x_4, x_5, x_6, x_242); +return x_249; } else { -uint8_t x_45; -lean_dec(x_17); -lean_dec(x_10); +uint8_t x_250; +lean_dec(x_26); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_45 = !lean_is_exclusive(x_18); -if (x_45 == 0) +lean_dec(x_3); +x_250 = !lean_is_exclusive(x_240); +if (x_250 == 0) { -return x_18; +return x_240; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_18, 0); -x_47 = lean_ctor_get(x_18, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_18); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_object* x_251; lean_object* x_252; lean_object* x_253; +x_251 = lean_ctor_get(x_240, 0); +x_252 = lean_ctor_get(x_240, 1); +lean_inc(x_252); +lean_inc(x_251); +lean_dec(x_240); +x_253 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_253, 0, x_251); +lean_ctor_set(x_253, 1, x_252); +return x_253; } } } else { -uint8_t x_49; -lean_dec(x_10); +uint8_t x_254; +lean_dec(x_26); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_49 = !lean_is_exclusive(x_12); -if (x_49 == 0) +x_254 = !lean_is_exclusive(x_236); +if (x_254 == 0) { -return x_12; +return x_236; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_12, 0); -x_51 = lean_ctor_get(x_12, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_12); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; -} +lean_object* x_255; lean_object* x_256; lean_object* x_257; +x_255 = lean_ctor_get(x_236, 0); +x_256 = lean_ctor_get(x_236, 1); +lean_inc(x_256); +lean_inc(x_255); +lean_dec(x_236); +x_257 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_257, 0, x_255); +lean_ctor_set(x_257, 1, x_256); +return x_257; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -lean_inc(x_1); -x_8 = l_Lean_MVarId_getTag(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +case 8: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +x_258 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_259 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_258, x_5, x_6, x_27); +x_260 = lean_ctor_get(x_259, 0); +lean_inc(x_260); +x_261 = lean_ctor_get(x_259, 1); +lean_inc(x_261); +lean_dec(x_259); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_11 = l_Lean_FVarId_getType(x_2, x_3, x_4, x_5, x_6, x_10); -if (lean_obj_tag(x_11) == 0) +x_262 = l_Lean_MVarId_assert(x_1, x_260, x_15, x_2, x_3, x_4, x_5, x_6, x_261); +if (lean_obj_tag(x_262) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); +lean_object* x_263; lean_object* x_264; uint8_t x_265; lean_object* x_266; +x_263 = lean_ctor_get(x_262, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_262, 1); +lean_inc(x_264); +lean_dec(x_262); +x_265 = 0; lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_14 = lean_whnf(x_12, x_3, x_4, x_5, x_6, x_13); -if (lean_obj_tag(x_14) == 0) +x_266 = l_Lean_Meta_intro1Core(x_263, x_265, x_3, x_4, x_5, x_6, x_264); +if (lean_obj_tag(x_266) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_Expr_getAppFn(x_15); -if (lean_obj_tag(x_17) == 4) +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; +x_267 = lean_ctor_get(x_266, 0); +lean_inc(x_267); +x_268 = lean_ctor_get(x_266, 1); +lean_inc(x_268); +lean_dec(x_266); +x_269 = lean_ctor_get(x_267, 0); +lean_inc(x_269); +x_270 = lean_ctor_get(x_267, 1); +lean_inc(x_270); +lean_dec(x_267); +x_271 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_272 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_273 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_270); +x_274 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__8___boxed), 12, 7); +lean_closure_set(x_274, 0, x_270); +lean_closure_set(x_274, 1, x_271); +lean_closure_set(x_274, 2, x_269); +lean_closure_set(x_274, 3, x_26); +lean_closure_set(x_274, 4, x_272); +lean_closure_set(x_274, 5, x_273); +lean_closure_set(x_274, 6, x_9); +x_275 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_270, x_274, x_3, x_4, x_5, x_6, x_268); +return x_275; +} +else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -lean_inc(x_18); -x_19 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_18, x_3, x_4, x_5, x_6, x_16); -if (lean_obj_tag(x_19) == 0) +uint8_t x_276; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_276 = !lean_is_exclusive(x_266); +if (x_276 == 0) { -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -if (lean_obj_tag(x_20) == 5) +return x_266; +} +else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_dec(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_casesOnSuffix; -x_23 = l_Lean_Name_str___override(x_18, x_22); -x_24 = lean_box(0); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_25 = l_Lean_Meta_mkRecursorInfo(x_23, x_24, x_3, x_4, x_5, x_6, x_21); -if (lean_obj_tag(x_25) == 0) +lean_object* x_277; lean_object* x_278; lean_object* x_279; +x_277 = lean_ctor_get(x_266, 0); +x_278 = lean_ctor_get(x_266, 1); +lean_inc(x_278); +lean_inc(x_277); +lean_dec(x_266); +x_279 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_279, 0, x_277); +lean_ctor_set(x_279, 1, x_278); +return x_279; +} +} +} +else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = l_Lean_Meta_RecursorInfo_numIndices(x_26); -x_29 = lean_unsigned_to_nat(0u); -x_30 = lean_nat_dec_lt(x_29, x_28); -lean_dec(x_28); -if (x_30 == 0) +uint8_t x_280; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_280 = !lean_is_exclusive(x_262); +if (x_280 == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; -lean_inc(x_3); -lean_inc(x_26); -lean_inc(x_1); -x_32 = l_Lean_Meta_getMajorTypeIndices(x_1, x_31, x_26, x_15, x_3, x_4, x_5, x_6, x_27); -if (lean_obj_tag(x_32) == 0) +return x_262; +} +else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); +lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_281 = lean_ctor_get(x_262, 0); +x_282 = lean_ctor_get(x_262, 1); +lean_inc(x_282); +lean_inc(x_281); +lean_dec(x_262); +x_283 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_283, 0, x_281); +lean_ctor_set(x_283, 1, x_282); +return x_283; +} +} +} +case 9: +{ +lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; +x_284 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_285 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_284, x_5, x_6, x_27); +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_285, 1); +lean_inc(x_287); +lean_dec(x_285); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_26); -lean_inc(x_2); -lean_inc(x_1); -x_35 = l_Lean_Meta_mkRecursorAppPrefix(x_1, x_31, x_2, x_26, x_33, x_3, x_4, x_5, x_6, x_34); -if (lean_obj_tag(x_35) == 0) +x_288 = l_Lean_MVarId_assert(x_1, x_286, x_15, x_2, x_3, x_4, x_5, x_6, x_287); +if (lean_obj_tag(x_288) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_mkAppN(x_36, x_33); -lean_dec(x_33); -lean_inc(x_2); -x_39 = l_Lean_Expr_fvar___override(x_2); -x_40 = l_Lean_Expr_app___override(x_38, x_39); +lean_object* x_289; lean_object* x_290; uint8_t x_291; lean_object* x_292; +x_289 = lean_ctor_get(x_288, 0); +lean_inc(x_289); +x_290 = lean_ctor_get(x_288, 1); +lean_inc(x_290); +lean_dec(x_288); +x_291 = 0; lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_40); -x_41 = lean_infer_type(x_40, x_3, x_4, x_5, x_6, x_37); -if (lean_obj_tag(x_41) == 0) +x_292 = l_Lean_Meta_intro1Core(x_289, x_291, x_3, x_4, x_5, x_6, x_290); +if (lean_obj_tag(x_292) == 0) +{ +lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_293 = lean_ctor_get(x_292, 0); +lean_inc(x_293); +x_294 = lean_ctor_get(x_292, 1); +lean_inc(x_294); +lean_dec(x_292); +x_295 = lean_ctor_get(x_293, 0); +lean_inc(x_295); +x_296 = lean_ctor_get(x_293, 1); +lean_inc(x_296); +lean_dec(x_293); +x_297 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_298 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_299 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_296); +x_300 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__9___boxed), 12, 7); +lean_closure_set(x_300, 0, x_296); +lean_closure_set(x_300, 1, x_297); +lean_closure_set(x_300, 2, x_295); +lean_closure_set(x_300, 3, x_26); +lean_closure_set(x_300, 4, x_298); +lean_closure_set(x_300, 5, x_299); +lean_closure_set(x_300, 6, x_9); +x_301 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_296, x_300, x_3, x_4, x_5, x_6, x_294); +return x_301; +} +else +{ +uint8_t x_302; +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_302 = !lean_is_exclusive(x_292); +if (x_302 == 0) +{ +return x_292; +} +else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_Meta_RecursorInfo_numMinors(x_26); -lean_dec(x_26); -x_45 = lean_unsigned_to_nat(1u); -x_46 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_46, 0, x_29); -lean_ctor_set(x_46, 1, x_44); -lean_ctor_set(x_46, 2, x_45); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_40); -lean_ctor_set(x_47, 1, x_42); -x_48 = l_Lean_Meta_Grind_cases___lambda__1___closed__1; -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_50 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(x_1, x_2, x_9, x_31, x_46, x_46, x_49, x_29, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_43); -lean_dec(x_46); -if (lean_obj_tag(x_50) == 0) +lean_object* x_303; lean_object* x_304; lean_object* x_305; +x_303 = lean_ctor_get(x_292, 0); +x_304 = lean_ctor_get(x_292, 1); +lean_inc(x_304); +lean_inc(x_303); +lean_dec(x_292); +x_305 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_305, 0, x_303); +lean_ctor_set(x_305, 1, x_304); +return x_305; +} +} +} +else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -x_53 = lean_ctor_get(x_50, 1); -lean_inc(x_53); -lean_dec(x_50); -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_ctor_get(x_52, 0); -lean_inc(x_55); -lean_dec(x_52); -x_56 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_55, x_3, x_4, x_5, x_6, x_53); +uint8_t x_306; +lean_dec(x_26); +lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) +x_306 = !lean_is_exclusive(x_288); +if (x_306 == 0) { -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_56, 0); -lean_dec(x_58); -x_59 = lean_array_to_list(x_54); -lean_ctor_set(x_56, 0, x_59); -return x_56; +return x_288; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_56, 1); -lean_inc(x_60); -lean_dec(x_56); -x_61 = lean_array_to_list(x_54); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_60); -return x_62; +lean_object* x_307; lean_object* x_308; lean_object* x_309; +x_307 = lean_ctor_get(x_288, 0); +x_308 = lean_ctor_get(x_288, 1); +lean_inc(x_308); +lean_inc(x_307); +lean_dec(x_288); +x_309 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_309, 0, x_307); +lean_ctor_set(x_309, 1, x_308); +return x_309; } } +} +case 10: +{ +lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; +x_310 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_311 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_310, x_5, x_6, x_27); +x_312 = lean_ctor_get(x_311, 0); +lean_inc(x_312); +x_313 = lean_ctor_get(x_311, 1); +lean_inc(x_313); +lean_dec(x_311); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_314 = l_Lean_MVarId_assert(x_1, x_312, x_15, x_2, x_3, x_4, x_5, x_6, x_313); +if (lean_obj_tag(x_314) == 0) +{ +lean_object* x_315; lean_object* x_316; uint8_t x_317; lean_object* x_318; +x_315 = lean_ctor_get(x_314, 0); +lean_inc(x_315); +x_316 = lean_ctor_get(x_314, 1); +lean_inc(x_316); +lean_dec(x_314); +x_317 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_318 = l_Lean_Meta_intro1Core(x_315, x_317, x_3, x_4, x_5, x_6, x_316); +if (lean_obj_tag(x_318) == 0) +{ +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; +x_319 = lean_ctor_get(x_318, 0); +lean_inc(x_319); +x_320 = lean_ctor_get(x_318, 1); +lean_inc(x_320); +lean_dec(x_318); +x_321 = lean_ctor_get(x_319, 0); +lean_inc(x_321); +x_322 = lean_ctor_get(x_319, 1); +lean_inc(x_322); +lean_dec(x_319); +x_323 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_324 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_325 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_322); +x_326 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__10___boxed), 12, 7); +lean_closure_set(x_326, 0, x_322); +lean_closure_set(x_326, 1, x_323); +lean_closure_set(x_326, 2, x_321); +lean_closure_set(x_326, 3, x_26); +lean_closure_set(x_326, 4, x_324); +lean_closure_set(x_326, 5, x_325); +lean_closure_set(x_326, 6, x_9); +x_327 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_322, x_326, x_3, x_4, x_5, x_6, x_320); +return x_327; +} else { -uint8_t x_63; +uint8_t x_328; +lean_dec(x_26); +lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -x_63 = !lean_is_exclusive(x_50); -if (x_63 == 0) +x_328 = !lean_is_exclusive(x_318); +if (x_328 == 0) { -return x_50; +return x_318; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_50, 0); -x_65 = lean_ctor_get(x_50, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_50); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_329; lean_object* x_330; lean_object* x_331; +x_329 = lean_ctor_get(x_318, 0); +x_330 = lean_ctor_get(x_318, 1); +lean_inc(x_330); +lean_inc(x_329); +lean_dec(x_318); +x_331 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_331, 0, x_329); +lean_ctor_set(x_331, 1, x_330); +return x_331; } } } else { -uint8_t x_67; -lean_dec(x_40); +uint8_t x_332; lean_dec(x_26); lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_67 = !lean_is_exclusive(x_41); -if (x_67 == 0) +x_332 = !lean_is_exclusive(x_314); +if (x_332 == 0) { -return x_41; +return x_314; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_41, 0); -x_69 = lean_ctor_get(x_41, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_41); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +lean_object* x_333; lean_object* x_334; lean_object* x_335; +x_333 = lean_ctor_get(x_314, 0); +x_334 = lean_ctor_get(x_314, 1); +lean_inc(x_334); +lean_inc(x_333); +lean_dec(x_314); +x_335 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_335, 0, x_333); +lean_ctor_set(x_335, 1, x_334); +return x_335; } } } +default: +{ +lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; +x_336 = l_Lean_Meta_Grind_cases___lambda__13___closed__2; +x_337 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_336, x_5, x_6, x_27); +x_338 = lean_ctor_get(x_337, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_337, 1); +lean_inc(x_339); +lean_dec(x_337); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_340 = l_Lean_MVarId_assert(x_1, x_338, x_15, x_2, x_3, x_4, x_5, x_6, x_339); +if (lean_obj_tag(x_340) == 0) +{ +lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_344; +x_341 = lean_ctor_get(x_340, 0); +lean_inc(x_341); +x_342 = lean_ctor_get(x_340, 1); +lean_inc(x_342); +lean_dec(x_340); +x_343 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_344 = l_Lean_Meta_intro1Core(x_341, x_343, x_3, x_4, x_5, x_6, x_342); +if (lean_obj_tag(x_344) == 0) +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; +x_345 = lean_ctor_get(x_344, 0); +lean_inc(x_345); +x_346 = lean_ctor_get(x_344, 1); +lean_inc(x_346); +lean_dec(x_344); +x_347 = lean_ctor_get(x_345, 0); +lean_inc(x_347); +x_348 = lean_ctor_get(x_345, 1); +lean_inc(x_348); +lean_dec(x_345); +x_349 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +x_350 = l_Lean_Meta_Grind_cases___lambda__13___closed__4; +x_351 = l_Lean_Meta_Grind_cases___lambda__12___closed__1; +lean_inc(x_348); +x_352 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__11___boxed), 12, 7); +lean_closure_set(x_352, 0, x_348); +lean_closure_set(x_352, 1, x_349); +lean_closure_set(x_352, 2, x_347); +lean_closure_set(x_352, 3, x_26); +lean_closure_set(x_352, 4, x_350); +lean_closure_set(x_352, 5, x_351); +lean_closure_set(x_352, 6, x_9); +x_353 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_348, x_352, x_3, x_4, x_5, x_6, x_346); +return x_353; +} else { -uint8_t x_71; -lean_dec(x_33); +uint8_t x_354; lean_dec(x_26); lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_71 = !lean_is_exclusive(x_35); -if (x_71 == 0) +x_354 = !lean_is_exclusive(x_344); +if (x_354 == 0) { -return x_35; +return x_344; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_35, 0); -x_73 = lean_ctor_get(x_35, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_35); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; +lean_object* x_355; lean_object* x_356; lean_object* x_357; +x_355 = lean_ctor_get(x_344, 0); +x_356 = lean_ctor_get(x_344, 1); +lean_inc(x_356); +lean_inc(x_355); +lean_dec(x_344); +x_357 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_357, 0, x_355); +lean_ctor_set(x_357, 1, x_356); +return x_357; } } } else { -uint8_t x_75; +uint8_t x_358; lean_dec(x_26); lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_75 = !lean_is_exclusive(x_32); -if (x_75 == 0) +x_358 = !lean_is_exclusive(x_340); +if (x_358 == 0) { -return x_32; +return x_340; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_32, 0); -x_77 = lean_ctor_get(x_32, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_32); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -return x_78; +lean_object* x_359; lean_object* x_360; lean_object* x_361; +x_359 = lean_ctor_get(x_340, 0); +x_360 = lean_ctor_get(x_340, 1); +lean_inc(x_360); +lean_inc(x_359); +lean_dec(x_340); +x_361 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_361, 0, x_359); +lean_ctor_set(x_361, 1, x_360); +return x_361; +} +} } } } else { -lean_object* x_79; +lean_object* x_362; lean_dec(x_15); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_79 = l_Lean_Meta_generalizeIndices(x_1, x_2, x_3, x_4, x_5, x_6, x_27); -if (lean_obj_tag(x_79) == 0) +x_362 = l_Lean_Meta_generalizeIndices_x27(x_1, x_2, x_24, x_3, x_4, x_5, x_6, x_27); +if (lean_obj_tag(x_362) == 0) { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; size_t x_85; size_t x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -lean_dec(x_79); -x_82 = lean_ctor_get(x_80, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_80, 2); -lean_inc(x_83); -x_84 = lean_ctor_get(x_80, 1); -lean_inc(x_84); -lean_dec(x_80); -x_85 = lean_array_size(x_84); -x_86 = 0; -x_87 = l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(x_85, x_86, x_84); -x_88 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; -lean_inc(x_82); -x_89 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__1___boxed), 11, 6); -lean_closure_set(x_89, 0, x_82); -lean_closure_set(x_89, 1, x_88); -lean_closure_set(x_89, 2, x_83); -lean_closure_set(x_89, 3, x_26); -lean_closure_set(x_89, 4, x_87); -lean_closure_set(x_89, 5, x_9); -x_90 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_82, x_89, x_3, x_4, x_5, x_6, x_81); -return x_90; +lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; size_t x_368; size_t x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; +x_363 = lean_ctor_get(x_362, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_362, 1); +lean_inc(x_364); +lean_dec(x_362); +x_365 = lean_ctor_get(x_363, 0); +lean_inc(x_365); +x_366 = lean_ctor_get(x_363, 2); +lean_inc(x_366); +x_367 = lean_ctor_get(x_363, 1); +lean_inc(x_367); +lean_dec(x_363); +x_368 = lean_array_size(x_367); +x_369 = 0; +lean_inc(x_367); +x_370 = l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(x_368, x_369, x_367); +x_371 = l_Lean_Meta_Grind_cases_throwInductiveExpected___closed__3; +lean_inc(x_365); +x_372 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__12___boxed), 12, 7); +lean_closure_set(x_372, 0, x_365); +lean_closure_set(x_372, 1, x_371); +lean_closure_set(x_372, 2, x_366); +lean_closure_set(x_372, 3, x_26); +lean_closure_set(x_372, 4, x_370); +lean_closure_set(x_372, 5, x_9); +lean_closure_set(x_372, 6, x_367); +x_373 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_365, x_372, x_3, x_4, x_5, x_6, x_364); +return x_373; } else { -uint8_t x_91; +uint8_t x_374; lean_dec(x_26); lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_91 = !lean_is_exclusive(x_79); -if (x_91 == 0) +x_374 = !lean_is_exclusive(x_362); +if (x_374 == 0) { -return x_79; +return x_362; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_79, 0); -x_93 = lean_ctor_get(x_79, 1); -lean_inc(x_93); -lean_inc(x_92); -lean_dec(x_79); -x_94 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_94, 0, x_92); -lean_ctor_set(x_94, 1, x_93); -return x_94; +lean_object* x_375; lean_object* x_376; lean_object* x_377; +x_375 = lean_ctor_get(x_362, 0); +x_376 = lean_ctor_get(x_362, 1); +lean_inc(x_376); +lean_inc(x_375); +lean_dec(x_362); +x_377 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_377, 0, x_375); +lean_ctor_set(x_377, 1, x_376); +return x_377; } } } } else { -uint8_t x_95; +uint8_t x_378; lean_dec(x_15); lean_dec(x_9); lean_dec(x_6); @@ -1216,46 +11108,46 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_95 = !lean_is_exclusive(x_25); -if (x_95 == 0) +x_378 = !lean_is_exclusive(x_25); +if (x_378 == 0) { return x_25; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_25, 0); -x_97 = lean_ctor_get(x_25, 1); -lean_inc(x_97); -lean_inc(x_96); +lean_object* x_379; lean_object* x_380; lean_object* x_381; +x_379 = lean_ctor_get(x_25, 0); +x_380 = lean_ctor_get(x_25, 1); +lean_inc(x_380); +lean_inc(x_379); lean_dec(x_25); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -return x_98; +x_381 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_381, 0, x_379); +lean_ctor_set(x_381, 1, x_380); +return x_381; } } } else { -lean_object* x_99; lean_object* x_100; +lean_object* x_382; lean_object* x_383; lean_dec(x_20); lean_dec(x_18); lean_dec(x_9); -x_99 = lean_ctor_get(x_19, 1); -lean_inc(x_99); +x_382 = lean_ctor_get(x_19, 1); +lean_inc(x_382); lean_dec(x_19); -x_100 = l_Lean_Meta_Grind_cases_throwInductiveExpected(x_1, x_2, lean_box(0), x_15, x_3, x_4, x_5, x_6, x_99); +x_383 = l_Lean_Meta_Grind_cases_throwInductiveExpected(x_1, x_2, lean_box(0), x_15, x_3, x_4, x_5, x_6, x_382); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_100; +return x_383; } } else { -uint8_t x_101; +uint8_t x_384; lean_dec(x_18); lean_dec(x_15); lean_dec(x_9); @@ -1265,42 +11157,42 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_101 = !lean_is_exclusive(x_19); -if (x_101 == 0) +x_384 = !lean_is_exclusive(x_19); +if (x_384 == 0) { return x_19; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_19, 0); -x_103 = lean_ctor_get(x_19, 1); -lean_inc(x_103); -lean_inc(x_102); +lean_object* x_385; lean_object* x_386; lean_object* x_387; +x_385 = lean_ctor_get(x_19, 0); +x_386 = lean_ctor_get(x_19, 1); +lean_inc(x_386); +lean_inc(x_385); lean_dec(x_19); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -return x_104; +x_387 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_387, 0, x_385); +lean_ctor_set(x_387, 1, x_386); +return x_387; } } } else { -lean_object* x_105; +lean_object* x_388; lean_dec(x_17); lean_dec(x_9); -x_105 = l_Lean_Meta_Grind_cases_throwInductiveExpected(x_1, x_2, lean_box(0), x_15, x_3, x_4, x_5, x_6, x_16); +x_388 = l_Lean_Meta_Grind_cases_throwInductiveExpected(x_1, x_2, lean_box(0), x_15, x_3, x_4, x_5, x_6, x_16); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_105; +return x_388; } } else { -uint8_t x_106; +uint8_t x_389; lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); @@ -1308,29 +11200,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_106 = !lean_is_exclusive(x_14); -if (x_106 == 0) +x_389 = !lean_is_exclusive(x_14); +if (x_389 == 0) { return x_14; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_14, 0); -x_108 = lean_ctor_get(x_14, 1); -lean_inc(x_108); -lean_inc(x_107); +lean_object* x_390; lean_object* x_391; lean_object* x_392; +x_390 = lean_ctor_get(x_14, 0); +x_391 = lean_ctor_get(x_14, 1); +lean_inc(x_391); +lean_inc(x_390); lean_dec(x_14); -x_109 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -return x_109; +x_392 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_392, 0, x_390); +lean_ctor_set(x_392, 1, x_391); +return x_392; } } } else { -uint8_t x_110; +uint8_t x_393; lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); @@ -1338,52 +11230,52 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_110 = !lean_is_exclusive(x_11); -if (x_110 == 0) +x_393 = !lean_is_exclusive(x_11); +if (x_393 == 0) { return x_11; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_11, 0); -x_112 = lean_ctor_get(x_11, 1); -lean_inc(x_112); -lean_inc(x_111); +lean_object* x_394; lean_object* x_395; lean_object* x_396; +x_394 = lean_ctor_get(x_11, 0); +x_395 = lean_ctor_get(x_11, 1); +lean_inc(x_395); +lean_inc(x_394); lean_dec(x_11); -x_113 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -return x_113; +x_396 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_396, 0, x_394); +lean_ctor_set(x_396, 1, x_395); +return x_396; } } } else { -uint8_t x_114; +uint8_t x_397; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_114 = !lean_is_exclusive(x_8); -if (x_114 == 0) +x_397 = !lean_is_exclusive(x_8); +if (x_397 == 0) { return x_8; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_8, 0); -x_116 = lean_ctor_get(x_8, 1); -lean_inc(x_116); -lean_inc(x_115); +lean_object* x_398; lean_object* x_399; lean_object* x_400; +x_398 = lean_ctor_get(x_8, 0); +x_399 = lean_ctor_get(x_8, 1); +lean_inc(x_399); +lean_inc(x_398); lean_dec(x_8); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +x_400 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_400, 0, x_398); +lean_ctor_set(x_400, 1, x_399); +return x_400; } } } @@ -1393,52 +11285,509 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases(lean_object* x_1, lean_object* { lean_object* x_8; lean_object* x_9; lean_inc(x_1); -x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__2), 7, 2); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_cases___lambda__13), 7, 2); lean_closure_set(x_8, 0, x_1); lean_closure_set(x_8, 1, x_2); x_9 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); return x_9; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -lean_object* x_10; -x_10 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +lean_object* x_19; +x_19 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_5); -return x_10; +return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_16; -x_16 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_6); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_5); -return x_16; +return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_15; -x_15 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_5); -lean_dec(x_4); -return x_15; +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_cases___lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_cases___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_cases___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_5); -return x_12; +return x_13; } } lean_object* initialize_Lean_Meta_Tactic_Cases(uint8_t builtin, lean_object*); @@ -1472,8 +11821,15 @@ l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__3 = _i lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__3); l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4 = _init_l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4(); lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_cases___spec__1___closed__4); -l_Lean_Meta_Grind_cases___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_cases___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__1___closed__1); +l_Lean_Meta_Grind_cases___lambda__12___closed__1 = _init_l_Lean_Meta_Grind_cases___lambda__12___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__12___closed__1); +l_Lean_Meta_Grind_cases___lambda__13___closed__1 = _init_l_Lean_Meta_Grind_cases___lambda__13___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__13___closed__1); +l_Lean_Meta_Grind_cases___lambda__13___closed__2 = _init_l_Lean_Meta_Grind_cases___lambda__13___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__13___closed__2); +l_Lean_Meta_Grind_cases___lambda__13___closed__3 = _init_l_Lean_Meta_Grind_cases___lambda__13___closed__3(); +l_Lean_Meta_Grind_cases___lambda__13___closed__4 = _init_l_Lean_Meta_Grind_cases___lambda__13___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_cases___lambda__13___closed__4); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/CasesMatch.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/CasesMatch.c new file mode 100644 index 000000000000..eed01a9a4bfc --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/CasesMatch.c @@ -0,0 +1,1154 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.CasesMatch +// Imports: Lean.Meta.Tactic.Util Lean.Meta.Tactic.Cases Lean.Meta.Match.MatcherApp +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_updateTags(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_sort___override(lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2; +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1; +lean_object* lean_array_fget(lean_object*, lean_object*); +lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_matchMatcherApp_x3f___at___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_generalizeMatchDiscrs_mkNewTarget___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6; +lean_object* l_Lean_stringToMessageData(lean_object*); +lean_object* l_Lean_Level_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_updateTags___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5; +lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_to_list(lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_casesMatch___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_forallMetaBoundedTelescope(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); +lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_withNewEqs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_get_match_equations_for(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4; +lean_object* l_Lean_MVarId_setTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2; +lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* l_Lean_Expr_mvarId_x21(lean_object*); +static lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_MVarId_getType(x_1, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = 0; +x_14 = 1; +x_15 = 1; +x_16 = l_Lean_Meta_mkForallFVars(x_3, x_11, x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Meta_mkLambdaFVars(x_2, x_17, x_13, x_14, x_13, x_15, x_5, x_6, x_7, x_8, x_18); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_4); +lean_ctor_set(x_19, 0, x_22); +return x_19; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_19, 0); +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_19); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_4); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +return x_26; +} +} +else +{ +uint8_t x_27; +lean_dec(x_4); +x_27 = !lean_is_exclusive(x_19); +if (x_27 == 0) +{ +return x_19; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_19, 0); +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_19); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +else +{ +uint8_t x_31; +lean_dec(x_4); +x_31 = !lean_is_exclusive(x_16); +if (x_31 == 0) +{ +return x_16; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_16, 0); +x_33 = lean_ctor_get(x_16, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_16); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_4); +x_35 = !lean_is_exclusive(x_10); +if (x_35 == 0) +{ +return x_10; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_10, 0); +x_37 = lean_ctor_get(x_10, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_10); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +lean_inc(x_3); +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1___boxed), 9, 2); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_3); +x_11 = l_Lean_Meta_withNewEqs___rarg(x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Level_ofNat(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = l_Lean_Expr_getAppFn(x_2); +x_10 = lean_ctor_get(x_3, 4); +lean_inc(x_10); +x_11 = l_Lean_mkAppN(x_9, x_10); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2; +x_13 = l_Lean_Expr_app___override(x_11, x_12); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_14 = lean_infer_type(x_13, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_ctor_get(x_3, 6); +lean_inc(x_17); +lean_dec(x_3); +x_18 = lean_array_get_size(x_17); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2___boxed), 9, 2); +lean_closure_set(x_20, 0, x_1); +lean_closure_set(x_20, 1, x_17); +x_21 = 0; +x_22 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_15, x_19, x_20, x_21, x_4, x_5, x_6, x_7, x_16); +return x_22; +} +else +{ +uint8_t x_23; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_14); +if (x_23 == 0) +{ +return x_14; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_14, 0); +x_25 = lean_ctor_get(x_14, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_14); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_6, x_5); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_2); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; +x_15 = lean_array_uget(x_4, x_6); +x_16 = l_Lean_Expr_mvarId_x21(x_15); +lean_dec(x_15); +lean_inc(x_7); +lean_inc(x_2); +x_17 = l_Lean_Name_num___override(x_2, x_7); +x_18 = l_Lean_MVarId_setTag(x_16, x_17, x_8, x_9, x_10, x_11, x_12); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_add(x_7, x_20); +lean_dec(x_7); +x_22 = 1; +x_23 = lean_usize_add(x_6, x_22); +x_6 = x_23; +x_7 = x_21; +x_12 = x_19; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_updateTags(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_MVarId_getTag(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_array_get_size(x_2); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_dec_eq(x_11, x_12); +if (x_13 == 0) +{ +lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_11); +x_14 = lean_box(0); +x_15 = lean_array_size(x_2); +x_16 = 0; +x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1(x_2, x_9, x_14, x_2, x_15, x_16, x_12, x_3, x_4, x_5, x_6, x_10); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_17, 0, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; +} +} +else +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_nat_dec_lt(x_24, x_11); +lean_dec(x_11); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = l_Lean_instInhabitedExpr; +x_27 = l_outOfBounds___rarg(x_26); +x_28 = l_Lean_Expr_mvarId_x21(x_27); +lean_dec(x_27); +x_29 = l_Lean_MVarId_setTag(x_28, x_9, x_3, x_4, x_5, x_6, x_10); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_array_fget(x_2, x_24); +x_31 = l_Lean_Expr_mvarId_x21(x_30); +lean_dec(x_30); +x_32 = l_Lean_MVarId_setTag(x_31, x_9, x_3, x_4, x_5, x_6, x_10); +return x_32; +} +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_8); +if (x_33 == 0) +{ +return x_8; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_8, 0); +x_35 = lean_ctor_get(x_8, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_8); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_casesMatch_updateTags___spec__1(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch_updateTags___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Grind_casesMatch_updateTags(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_casesMatch___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Lean_Expr_mvarId_x21(x_5); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_Expr_mvarId_x21(x_9); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_13 = lean_get_match_equations_for(x_12, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_array_to_list(x_5); +x_18 = l_Lean_Expr_const___override(x_16, x_17); +x_19 = lean_ctor_get(x_1, 4); +lean_inc(x_19); +x_20 = l_Lean_mkAppN(x_18, x_19); +lean_dec(x_19); +x_21 = l_Lean_Expr_app___override(x_20, x_2); +x_22 = lean_ctor_get(x_1, 6); +lean_inc(x_22); +x_23 = l_Lean_mkAppN(x_21, x_22); +lean_dec(x_22); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_23); +x_24 = lean_infer_type(x_23, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_ctor_get(x_1, 8); +lean_inc(x_27); +lean_dec(x_1); +x_28 = lean_array_get_size(x_27); +lean_dec(x_27); +x_29 = 2; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_30 = l_Lean_Meta_forallMetaBoundedTelescope(x_25, x_28, x_29, x_7, x_8, x_9, x_10, x_26); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_mkAppN(x_23, x_33); +x_35 = l_Lean_mkAppN(x_34, x_3); +lean_inc(x_4); +x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_4, x_35, x_7, x_8, x_9, x_10, x_32); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_Meta_Grind_casesMatch_updateTags(x_4, x_33, x_7, x_8, x_9, x_10, x_37); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +if (lean_obj_tag(x_38) == 0) +{ +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_38, 0); +lean_dec(x_40); +x_41 = lean_array_to_list(x_33); +x_42 = lean_box(0); +x_43 = l_List_mapTR_loop___at_Lean_Meta_Grind_casesMatch___spec__1(x_41, x_42); +lean_ctor_set(x_38, 0, x_43); +return x_38; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_44 = lean_ctor_get(x_38, 1); +lean_inc(x_44); +lean_dec(x_38); +x_45 = lean_array_to_list(x_33); +x_46 = lean_box(0); +x_47 = l_List_mapTR_loop___at_Lean_Meta_Grind_casesMatch___spec__1(x_45, x_46); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_44); +return x_48; +} +} +else +{ +uint8_t x_49; +lean_dec(x_33); +x_49 = !lean_is_exclusive(x_38); +if (x_49 == 0) +{ +return x_38; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_38, 0); +x_51 = lean_ctor_get(x_38, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_38); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +} +else +{ +uint8_t x_53; +lean_dec(x_23); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_53 = !lean_is_exclusive(x_30); +if (x_53 == 0) +{ +return x_30; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_30, 0); +x_55 = lean_ctor_get(x_30, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_30); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; +} +} +} +else +{ +uint8_t x_57; +lean_dec(x_23); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_24); +if (x_57 == 0) +{ +return x_24; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_24, 0); +x_59 = lean_ctor_get(x_24, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_24); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +uint8_t x_61; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_61 = !lean_is_exclusive(x_13); +if (x_61 == 0) +{ +return x_13; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_13, 0); +x_63 = lean_ctor_get(x_13, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_13); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("casesMatch", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1; +x_2 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`match`-expression expected", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = 0; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +x_9 = l_Lean_Meta_matchMatcherApp_x3f___at___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_generalizeMatchDiscrs_mkNewTarget___spec__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_indentExpr(x_1); +x_13 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3; +x_19 = l_Lean_Meta_throwTacticEx___rarg(x_18, x_2, x_17, x_3, x_4, x_5, x_6, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_9, 1); +lean_inc(x_20); +lean_dec(x_9); +x_21 = lean_ctor_get(x_10, 0); +lean_inc(x_21); +lean_dec(x_10); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_21); +lean_inc(x_2); +x_22 = l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls(x_2, x_1, x_21, x_3, x_4, x_5, x_6, x_20); +lean_dec(x_1); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +lean_inc(x_2); +x_27 = l_Lean_MVarId_getType(x_2, x_3, x_4, x_5, x_6, x_24); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; +x_28 = lean_ctor_get(x_21, 2); +lean_inc(x_28); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +x_31 = lean_box(0); +x_32 = l_Lean_Meta_Grind_casesMatch___lambda__1(x_21, x_25, x_26, x_2, x_30, x_31, x_3, x_4, x_5, x_6, x_29); +lean_dec(x_26); +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = lean_ctor_get(x_27, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_dec(x_27); +x_35 = lean_ctor_get(x_21, 1); +lean_inc(x_35); +x_36 = lean_ctor_get(x_28, 0); +lean_inc(x_36); +lean_dec(x_28); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_37 = l_Lean_Meta_getLevel(x_33, x_3, x_4, x_5, x_6, x_34); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_array_set(x_35, x_36, x_38); +lean_dec(x_36); +x_41 = lean_box(0); +x_42 = l_Lean_Meta_Grind_casesMatch___lambda__1(x_21, x_25, x_26, x_2, x_40, x_41, x_3, x_4, x_5, x_6, x_39); +lean_dec(x_26); +return x_42; +} +else +{ +uint8_t x_43; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_43 = !lean_is_exclusive(x_37); +if (x_43 == 0) +{ +return x_37; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_37, 0); +x_45 = lean_ctor_get(x_37, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_37); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_27); +if (x_47 == 0) +{ +return x_27; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_27, 0); +x_49 = lean_ctor_get(x_27, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_27); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_21); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_51 = !lean_is_exclusive(x_22); +if (x_51 == 0) +{ +return x_22; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_22, 0); +x_53 = lean_ctor_get(x_22, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_22); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +uint8_t x_55; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_9); +if (x_55 == 0) +{ +return x_9; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_9, 0); +x_57 = lean_ctor_get(x_9, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_9); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +lean_inc(x_1); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_casesMatch___lambda__2), 7, 2); +lean_closure_set(x_8, 0, x_2); +lean_closure_set(x_8, 1, x_1); +x_9 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_casesMatch___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_casesMatch___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_3); +return x_12; +} +} +lean_object* initialize_Lean_Meta_Tactic_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Cases(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Match_MatcherApp(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_CasesMatch(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Meta_Tactic_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Cases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Match_MatcherApp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1 = _init_l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__1); +l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2 = _init_l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch_mkMotiveAndRefls___closed__2); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__1); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__2); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__3); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__4); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__5); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__6); +l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_casesMatch___lambda__2___closed__7); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Combinators.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Combinators.c new file mode 100644 index 000000000000..76406a632a33 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Combinators.c @@ -0,0 +1,1388 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.Combinators +// Imports: Lean.Meta.Tactic.Grind.Types +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_toGrindTactic___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_toGrindTactic___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instOrElseGrindTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_iterate_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_x27_toGrindTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_andThen(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_try(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyToAll_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_orElse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyToAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instAndThenGrindTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_try(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_2); +x_11 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_2); +lean_ctor_set(x_16, 1, x_15); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_11, 0, x_17); +return x_11; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_dec(x_11); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_2); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; +} +} +else +{ +uint8_t x_23; +lean_dec(x_2); +x_23 = !lean_is_exclusive(x_11); +if (x_23 == 0) +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_11, 0); +lean_dec(x_24); +x_25 = !lean_is_exclusive(x_12); +if (x_25 == 0) +{ +return x_11; +} +else +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_12, 0); +lean_inc(x_26); +lean_dec(x_12); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_11, 0, x_27); +return x_11; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +lean_dec(x_11); +x_29 = lean_ctor_get(x_12, 0); +lean_inc(x_29); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + x_30 = x_12; +} else { + lean_dec_ref(x_12); + x_30 = lean_box(0); +} +if (lean_is_scalar(x_30)) { + x_31 = lean_alloc_ctor(1, 1, 0); +} else { + x_31 = x_30; +} +lean_ctor_set(x_31, 0, x_29); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_28); +return x_32; +} +} +} +else +{ +uint8_t x_33; +lean_dec(x_2); +x_33 = !lean_is_exclusive(x_11); +if (x_33 == 0) +{ +return x_11; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_11, 0); +x_35 = lean_ctor_get(x_11, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_11); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyToAll_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = l_List_reverse___rarg(x_3); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_2); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_2, 0); +x_16 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_15); +x_17 = lean_apply_9(x_1, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_16; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_10 = x_19; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_11 = _tmp_10; +} +goto _start; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_free_object(x_2); +lean_dec(x_15); +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_ctor_get(x_18, 0); +lean_inc(x_22); +lean_dec(x_18); +x_23 = l_List_appendTR___rarg(x_22, x_3); +x_2 = x_16; +x_3 = x_23; +x_11 = x_21; +goto _start; +} +} +else +{ +uint8_t x_25; +lean_free_object(x_2); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) +{ +return x_17; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_17, 0); +x_27 = lean_ctor_get(x_17, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_2, 0); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_2); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_29); +x_31 = lean_apply_9(x_1, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_29); +lean_ctor_set(x_34, 1, x_3); +x_2 = x_30; +x_3 = x_34; +x_11 = x_33; +goto _start; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_29); +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); +lean_dec(x_31); +x_37 = lean_ctor_get(x_32, 0); +lean_inc(x_37); +lean_dec(x_32); +x_38 = l_List_appendTR___rarg(x_37, x_3); +x_2 = x_30; +x_3 = x_38; +x_11 = x_36; +goto _start; +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_40 = lean_ctor_get(x_31, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_31, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_42 = x_31; +} else { + lean_dec_ref(x_31); + x_42 = lean_box(0); +} +if (lean_is_scalar(x_42)) { + x_43 = lean_alloc_ctor(1, 2, 0); +} else { + x_43 = x_42; +} +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_41); +return x_43; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyToAll(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_Meta_Grind_applyToAll_go(x_1, x_2, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_andThen(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_12 = lean_apply_9(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_12, 0, x_16); +return x_12; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +else +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = !lean_is_exclusive(x_13); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_13, 0); +x_23 = l_Lean_Meta_Grind_applyToAll(x_2, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_20); +if (lean_obj_tag(x_23) == 0) +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_23, 0); +lean_ctor_set(x_13, 0, x_25); +lean_ctor_set(x_23, 0, x_13); +return x_23; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_23, 0); +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_23); +lean_ctor_set(x_13, 0, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_13); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +else +{ +uint8_t x_29; +lean_free_object(x_13); +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) +{ +return x_23; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_13, 0); +lean_inc(x_33); +lean_dec(x_13); +x_34 = l_Lean_Meta_Grind_applyToAll(x_2, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_20); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + lean_ctor_release(x_34, 1); + x_37 = x_34; +} else { + lean_dec_ref(x_34); + x_37 = lean_box(0); +} +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_35); +if (lean_is_scalar(x_37)) { + x_39 = lean_alloc_ctor(0, 2, 0); +} else { + x_39 = x_37; +} +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_34, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_34, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + lean_ctor_release(x_34, 1); + x_42 = x_34; +} else { + lean_dec_ref(x_34); + x_42 = lean_box(0); +} +if (lean_is_scalar(x_42)) { + x_43 = lean_alloc_ctor(1, 2, 0); +} else { + x_43 = x_42; +} +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_41); +return x_43; +} +} +} +} +else +{ +uint8_t x_44; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_44 = !lean_is_exclusive(x_12); +if (x_44 == 0) +{ +return x_12; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_12, 0); +x_46 = lean_ctor_get(x_12, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_12); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instAndThenGrindTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_box(0); +x_13 = lean_apply_1(x_2, x_12); +x_14 = l_Lean_Meta_Grind_GrindTactic_andThen(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_iterate_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_3); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_2); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_2, 0); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_14); +x_16 = lean_apply_9(x_1, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_15; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_10 = x_18; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_11 = _tmp_10; +} +goto _start; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_2); +lean_dec(x_14); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l_List_appendTR___rarg(x_21, x_15); +x_2 = x_22; +x_11 = x_20; +goto _start; +} +} +else +{ +uint8_t x_24; +lean_free_object(x_2); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_16); +if (x_24 == 0) +{ +return x_16; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_16, 0); +x_26 = lean_ctor_get(x_16, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_16); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_2, 0); +x_29 = lean_ctor_get(x_2, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_2); +lean_inc(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_28); +x_30 = lean_apply_9(x_1, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_28); +lean_ctor_set(x_33, 1, x_3); +x_2 = x_29; +x_3 = x_33; +x_11 = x_32; +goto _start; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_28); +x_35 = lean_ctor_get(x_30, 1); +lean_inc(x_35); +lean_dec(x_30); +x_36 = lean_ctor_get(x_31, 0); +lean_inc(x_36); +lean_dec(x_31); +x_37 = l_List_appendTR___rarg(x_36, x_29); +x_2 = x_37; +x_11 = x_35; +goto _start; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_39 = lean_ctor_get(x_30, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_30, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_41 = x_30; +} else { + lean_dec_ref(x_30); + x_41 = lean_box(0); +} +if (lean_is_scalar(x_41)) { + x_42 = lean_alloc_ctor(1, 2, 0); +} else { + x_42 = x_41; +} +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_11); +x_13 = l_Lean_Meta_Grind_GrindTactic_iterate_go(x_1, x_12, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_13); +if (x_21 == 0) +{ +return x_13; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_orElse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_12 = lean_apply_9(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_apply_9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_15; +} +else +{ +uint8_t x_16; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_16 = !lean_is_exclusive(x_12); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_12, 0); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) +{ +return x_12; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_13, 0); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_12, 0, x_20); +return x_12; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_ctor_get(x_13, 0); +lean_inc(x_22); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + x_23 = x_13; +} else { + lean_dec_ref(x_13); + x_23 = lean_box(0); +} +if (lean_is_scalar(x_23)) { + x_24 = lean_alloc_ctor(1, 1, 0); +} else { + x_24 = x_23; +} +lean_ctor_set(x_24, 0, x_22); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +return x_25; +} +} +} +else +{ +uint8_t x_26; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_12); +if (x_26 == 0) +{ +return x_12; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_12, 0); +x_28 = lean_ctor_get(x_12, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_12); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instOrElseGrindTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_box(0); +x_13 = lean_apply_1(x_2, x_12); +x_14 = l_Lean_Meta_Grind_GrindTactic_andThen(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_2); +x_11 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_2, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_13, 1); +x_16 = lean_st_ref_get(x_2, x_15); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_13, 1, x_18); +lean_ctor_set(x_16, 0, x_13); +return x_16; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_13, 1, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_st_ref_get(x_2, x_23); +lean_dec(x_2); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; +} else { + lean_dec_ref(x_24); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_25); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_27; +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_11); +if (x_30 == 0) +{ +return x_11; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_11, 0); +x_32 = lean_ctor_get(x_11, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_11); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); +return x_1; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_toGrindTactic___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_toGrindTactic___lambda__3___boxed), 9, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_toGrindTactic___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_toGrindTactic___lambda__1___boxed), 9, 1); +lean_closure_set(x_12, 0, x_2); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_toGrindTactic___lambda__2), 10, 1); +lean_closure_set(x_13, 0, x_1); +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = l_Lean_Meta_Grind_toGrindTactic___closed__1; +x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get_uint8(x_18, sizeof(void*)*19); +if (x_19 == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_18); +lean_ctor_set(x_23, 1, x_22); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_17, 0, x_24); +return x_17; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_dec(x_17); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_18); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_25); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_18); +x_30 = !lean_is_exclusive(x_17); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_17, 0); +lean_dec(x_31); +x_32 = l_Lean_Meta_Grind_toGrindTactic___closed__2; +lean_ctor_set(x_17, 0, x_32); +return x_17; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_17, 1); +lean_inc(x_33); +lean_dec(x_17); +x_34 = l_Lean_Meta_Grind_toGrindTactic___closed__2; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +} +else +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_17); +if (x_36 == 0) +{ +return x_17; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_17, 0); +x_38 = lean_ctor_get(x_17, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_17); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_toGrindTactic___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_toGrindTactic___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_toGrindTactic___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindTactic_x27_toGrindTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_11, 0, x_14); +return x_11; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_11, 0); +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Combinators(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_toGrindTactic___closed__1 = _init_l_Lean_Meta_Grind_toGrindTactic___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_toGrindTactic___closed__1); +l_Lean_Meta_Grind_toGrindTactic___closed__2 = _init_l_Lean_Meta_Grind_toGrindTactic___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_toGrindTactic___closed__2); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c index 80d6a7357d90..625647f73bd3 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Core.c @@ -70,6 +70,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_updateMT___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse___closed__4; +lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_goEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Level_ofNat(lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); @@ -102,6 +103,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tact LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2(lean_object*, lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_closeGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -120,7 +122,6 @@ lean_object* l_Lean_Meta_mkHEq(lean_object*, lean_object*, lean_object*, lean_ob lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); -lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_addCongrTable(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__15; @@ -146,6 +147,7 @@ lean_object* l_Lean_Meta_Grind_getParents(lean_object*, lean_object*, lean_objec static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore_processTodo___closed__1; +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__10; lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; @@ -155,6 +157,7 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_a static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addHypothesis(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -177,6 +180,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithTrueEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__4; +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_closeGoalWithValuesEq___closed__6; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -192,7 +196,6 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_a LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_invertTrans_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_add_go___closed__2; -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1011,6 +1014,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta lean_object* x_4; uint64_t x_5; size_t x_6; lean_object* x_7; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); +lean_inc(x_3); x_5 = l_Lean_Meta_Grind_congrHash(x_4, x_3); x_6 = lean_uint64_to_usize(x_5); x_7 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__2(x_1, x_2, x_6, x_3); @@ -1030,7 +1034,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; x_12 = lean_st_ref_take(x_3, x_11); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); @@ -1049,7 +1053,7 @@ x_19 = lean_ctor_get(x_13, 4); lean_inc(x_19); x_20 = lean_ctor_get(x_13, 5); lean_inc(x_20); -x_21 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); +x_21 = lean_ctor_get_uint8(x_13, sizeof(void*)*19); x_22 = lean_ctor_get(x_13, 6); lean_inc(x_22); x_23 = lean_ctor_get(x_13, 7); @@ -1066,105 +1070,130 @@ x_28 = lean_ctor_get(x_13, 12); lean_inc(x_28); x_29 = lean_ctor_get(x_13, 13); lean_inc(x_29); +x_30 = lean_ctor_get(x_13, 14); +lean_inc(x_30); +x_31 = lean_ctor_get(x_13, 15); +lean_inc(x_31); +x_32 = lean_ctor_get(x_13, 16); +lean_inc(x_32); +x_33 = lean_ctor_get(x_13, 17); +lean_inc(x_33); +x_34 = lean_ctor_get(x_13, 18); +lean_inc(x_34); lean_inc(x_13); -x_30 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__1(x_13, x_18, x_1); -x_31 = !lean_is_exclusive(x_13); -if (x_31 == 0) +x_35 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__1(x_13, x_18, x_1); +x_36 = !lean_is_exclusive(x_13); +if (x_36 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_32 = lean_ctor_get(x_13, 13); -lean_dec(x_32); -x_33 = lean_ctor_get(x_13, 12); -lean_dec(x_33); -x_34 = lean_ctor_get(x_13, 11); -lean_dec(x_34); -x_35 = lean_ctor_get(x_13, 10); -lean_dec(x_35); -x_36 = lean_ctor_get(x_13, 9); -lean_dec(x_36); -x_37 = lean_ctor_get(x_13, 8); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_37 = lean_ctor_get(x_13, 18); lean_dec(x_37); -x_38 = lean_ctor_get(x_13, 7); +x_38 = lean_ctor_get(x_13, 17); lean_dec(x_38); -x_39 = lean_ctor_get(x_13, 6); +x_39 = lean_ctor_get(x_13, 16); lean_dec(x_39); -x_40 = lean_ctor_get(x_13, 5); +x_40 = lean_ctor_get(x_13, 15); lean_dec(x_40); -x_41 = lean_ctor_get(x_13, 4); +x_41 = lean_ctor_get(x_13, 14); lean_dec(x_41); -x_42 = lean_ctor_get(x_13, 3); +x_42 = lean_ctor_get(x_13, 13); lean_dec(x_42); -x_43 = lean_ctor_get(x_13, 2); +x_43 = lean_ctor_get(x_13, 12); lean_dec(x_43); -x_44 = lean_ctor_get(x_13, 1); +x_44 = lean_ctor_get(x_13, 11); lean_dec(x_44); -x_45 = lean_ctor_get(x_13, 0); +x_45 = lean_ctor_get(x_13, 10); lean_dec(x_45); -lean_ctor_set(x_13, 3, x_30); -x_46 = lean_st_ref_set(x_3, x_13, x_14); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) -{ -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_46, 0); +x_46 = lean_ctor_get(x_13, 9); +lean_dec(x_46); +x_47 = lean_ctor_get(x_13, 8); +lean_dec(x_47); +x_48 = lean_ctor_get(x_13, 7); lean_dec(x_48); -x_49 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; -lean_ctor_set(x_46, 0, x_49); -return x_46; +x_49 = lean_ctor_get(x_13, 6); +lean_dec(x_49); +x_50 = lean_ctor_get(x_13, 5); +lean_dec(x_50); +x_51 = lean_ctor_get(x_13, 4); +lean_dec(x_51); +x_52 = lean_ctor_get(x_13, 3); +lean_dec(x_52); +x_53 = lean_ctor_get(x_13, 2); +lean_dec(x_53); +x_54 = lean_ctor_get(x_13, 1); +lean_dec(x_54); +x_55 = lean_ctor_get(x_13, 0); +lean_dec(x_55); +lean_ctor_set(x_13, 3, x_35); +x_56 = lean_st_ref_set(x_3, x_13, x_14); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 0); +lean_dec(x_58); +x_59 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +lean_ctor_set(x_56, 0, x_59); +return x_56; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_46, 1); -lean_inc(x_50); -lean_dec(x_46); -x_51 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_50); -return x_52; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_56, 1); +lean_inc(x_60); +lean_dec(x_56); +x_61 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_60); +return x_62; } } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_dec(x_13); -x_53 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_53, 0, x_15); -lean_ctor_set(x_53, 1, x_16); -lean_ctor_set(x_53, 2, x_17); -lean_ctor_set(x_53, 3, x_30); -lean_ctor_set(x_53, 4, x_19); -lean_ctor_set(x_53, 5, x_20); -lean_ctor_set(x_53, 6, x_22); -lean_ctor_set(x_53, 7, x_23); -lean_ctor_set(x_53, 8, x_24); -lean_ctor_set(x_53, 9, x_25); -lean_ctor_set(x_53, 10, x_26); -lean_ctor_set(x_53, 11, x_27); -lean_ctor_set(x_53, 12, x_28); -lean_ctor_set(x_53, 13, x_29); -lean_ctor_set_uint8(x_53, sizeof(void*)*14, x_21); -x_54 = lean_st_ref_set(x_3, x_53, x_14); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - x_56 = x_54; +x_63 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_63, 0, x_15); +lean_ctor_set(x_63, 1, x_16); +lean_ctor_set(x_63, 2, x_17); +lean_ctor_set(x_63, 3, x_35); +lean_ctor_set(x_63, 4, x_19); +lean_ctor_set(x_63, 5, x_20); +lean_ctor_set(x_63, 6, x_22); +lean_ctor_set(x_63, 7, x_23); +lean_ctor_set(x_63, 8, x_24); +lean_ctor_set(x_63, 9, x_25); +lean_ctor_set(x_63, 10, x_26); +lean_ctor_set(x_63, 11, x_27); +lean_ctor_set(x_63, 12, x_28); +lean_ctor_set(x_63, 13, x_29); +lean_ctor_set(x_63, 14, x_30); +lean_ctor_set(x_63, 15, x_31); +lean_ctor_set(x_63, 16, x_32); +lean_ctor_set(x_63, 17, x_33); +lean_ctor_set(x_63, 18, x_34); +lean_ctor_set_uint8(x_63, sizeof(void*)*19, x_21); +x_64 = lean_st_ref_set(x_3, x_63, x_14); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_66 = x_64; } else { - lean_dec_ref(x_54); - x_56 = lean_box(0); + lean_dec_ref(x_64); + x_66 = lean_box(0); } -x_57 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; -if (lean_is_scalar(x_56)) { - x_58 = lean_alloc_ctor(0, 2, 0); +x_67 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1___closed__1; +if (lean_is_scalar(x_66)) { + x_68 = lean_alloc_ctor(0, 2, 0); } else { - x_58 = x_56; + x_68 = x_66; } -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_55); -return x_58; +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_65); +return x_68; } } } @@ -1252,7 +1281,7 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_14 = lean_ctor_get(x_1, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_1, 1); @@ -1261,202 +1290,308 @@ x_16 = lean_ctor_get(x_1, 3); lean_inc(x_16); lean_dec(x_1); x_17 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +if (lean_obj_tag(x_17) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_19 = lean_ctor_get(x_17, 1); +lean_object* x_18; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +lean_dec(x_16); +lean_dec(x_15); +x_19 = !lean_is_exclusive(x_17); +if (x_19 == 0) +{ +lean_object* x_20; x_20 = lean_ctor_get(x_17, 0); lean_dec(x_20); -x_21 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; -x_22 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); -x_23 = lean_ctor_get(x_22, 0); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_18); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; uint8_t x_24; lean_object* x_25; uint8_t x_81; +lean_dec(x_18); +x_23 = lean_ctor_get(x_17, 1); lean_inc(x_23); -x_24 = lean_unbox(x_23); -lean_dec(x_23); +lean_dec(x_17); +x_81 = l_Lean_Expr_isApp(x_15); +if (x_81 == 0) +{ +x_24 = x_81; +x_25 = x_23; +goto block_80; +} +else +{ +lean_object* x_82; +lean_inc(x_15); +x_82 = l_Lean_Meta_Grind_isCongrRoot(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); +if (lean_obj_tag(x_82) == 0) +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +lean_dec(x_82); +x_85 = lean_unbox(x_83); +lean_dec(x_83); +x_24 = x_85; +x_25 = x_84; +goto block_80; +} +else +{ +uint8_t x_86; +lean_dec(x_16); +lean_dec(x_15); +x_86 = !lean_is_exclusive(x_82); +if (x_86 == 0) +{ +return x_82; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_82, 0); +x_88 = lean_ctor_get(x_82, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_82); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; +} +} +} +block_80: +{ if (x_24 == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_free_object(x_17); -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); +lean_object* x_26; +lean_dec(x_15); x_26 = lean_box(0); -x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); -lean_inc(x_30); -lean_dec(x_28); x_1 = x_16; -x_2 = x_30; -x_11 = x_29; +x_2 = x_26; +x_11 = x_25; goto _start; } else { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_22); -if (x_32 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; +x_29 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_unbox(x_30); +lean_dec(x_30); +if (x_31 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_33 = lean_ctor_get(x_22, 1); -x_34 = lean_ctor_get(x_22, 0); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_box(0); +x_34 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); lean_dec(x_34); -lean_inc(x_15); -x_35 = l_Lean_MessageData_ofExpr(x_15); -x_36 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; -lean_ctor_set_tag(x_22, 7); -lean_ctor_set(x_22, 1, x_35); -lean_ctor_set(x_22, 0, x_36); -x_37 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_37); -lean_ctor_set(x_17, 0, x_22); -x_38 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_21, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_33); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_40); -lean_dec(x_39); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = lean_ctor_get(x_42, 0); -lean_inc(x_44); -lean_dec(x_42); +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); x_1 = x_16; -x_2 = x_44; -x_11 = x_43; +x_2 = x_37; +x_11 = x_36; goto _start; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_46 = lean_ctor_get(x_22, 1); -lean_inc(x_46); -lean_dec(x_22); +uint8_t x_39; +x_39 = !lean_is_exclusive(x_29); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_29, 1); +x_41 = lean_ctor_get(x_29, 0); +lean_dec(x_41); +x_42 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_40); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); lean_inc(x_15); -x_47 = l_Lean_MessageData_ofExpr(x_15); -x_48 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -x_50 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_50); -lean_ctor_set(x_17, 0, x_49); -x_51 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_21, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_46); +x_44 = l_Lean_MessageData_ofExpr(x_15); +x_45 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; +lean_ctor_set_tag(x_29, 7); +lean_ctor_set(x_29, 1, x_44); +lean_ctor_set(x_29, 0, x_45); +x_46 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_29); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_28, x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_43); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_50); +lean_dec(x_49); x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); lean_dec(x_51); -x_54 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_53); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); lean_dec(x_52); -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -x_57 = lean_ctor_get(x_55, 0); -lean_inc(x_57); -lean_dec(x_55); x_1 = x_16; -x_2 = x_57; -x_11 = x_56; +x_2 = x_54; +x_11 = x_53; goto _start; } +else +{ +uint8_t x_56; +lean_free_object(x_29); +lean_dec(x_16); +lean_dec(x_15); +x_56 = !lean_is_exclusive(x_42); +if (x_56 == 0) +{ +return x_42; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_42, 0); +x_58 = lean_ctor_get(x_42, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_42); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_59 = lean_ctor_get(x_17, 1); -lean_inc(x_59); -lean_dec(x_17); -x_60 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; -x_61 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_59); -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_unbox(x_62); -lean_dec(x_62); -if (x_63 == 0) +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_29, 1); +lean_inc(x_60); +lean_dec(x_29); +x_61 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_60); +if (lean_obj_tag(x_61) == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_64 = lean_ctor_get(x_61, 1); -lean_inc(x_64); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); lean_dec(x_61); -x_65 = lean_box(0); -x_66 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_65, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_64); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_ctor_get(x_67, 0); +lean_inc(x_15); +x_63 = l_Lean_MessageData_ofExpr(x_15); +x_64 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +x_66 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_67 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_28, x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_62); +x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); -lean_dec(x_67); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_69, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_70); +lean_dec(x_69); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +x_74 = lean_ctor_get(x_72, 0); +lean_inc(x_74); +lean_dec(x_72); x_1 = x_16; -x_2 = x_69; -x_11 = x_68; +x_2 = x_74; +x_11 = x_73; goto _start; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_71 = lean_ctor_get(x_61, 1); -lean_inc(x_71); +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_16); +lean_dec(x_15); +x_76 = lean_ctor_get(x_61, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_61, 1); +lean_inc(x_77); if (lean_is_exclusive(x_61)) { lean_ctor_release(x_61, 0); lean_ctor_release(x_61, 1); - x_72 = x_61; + x_78 = x_61; } else { lean_dec_ref(x_61); - x_72 = lean_box(0); + x_78 = lean_box(0); } -lean_inc(x_15); -x_73 = l_Lean_MessageData_ofExpr(x_15); -x_74 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__6; -if (lean_is_scalar(x_72)) { - x_75 = lean_alloc_ctor(7, 2, 0); +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); } else { - x_75 = x_72; - lean_ctor_set_tag(x_75, 7); + x_79 = x_78; } -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_73); -x_76 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_77 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -x_78 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_60, x_77, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_71); -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec(x_78); -x_81 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___lambda__1(x_15, x_79, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_80); -lean_dec(x_79); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = lean_ctor_get(x_82, 0); -lean_inc(x_84); -lean_dec(x_82); -x_1 = x_16; -x_2 = x_84; -x_11 = x_83; -goto _start; +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +} +} +} +} +} +else +{ +uint8_t x_90; +lean_dec(x_16); +lean_dec(x_15); +x_90 = !lean_is_exclusive(x_17); +if (x_90 == 0) +{ +return x_17; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_17, 0); +x_92 = lean_ctor_get(x_17, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_17); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; } } } @@ -1465,7 +1600,7 @@ goto _start; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_11 = l_Lean_Meta_Grind_getParentsAndReset(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); @@ -1475,6 +1610,9 @@ lean_dec(x_11); x_14 = lean_box(0); lean_inc(x_12); x_15 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4(x_12, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) { @@ -1496,6 +1634,30 @@ lean_ctor_set(x_19, 1, x_18); return x_19; } } +else +{ +uint8_t x_20; +lean_dec(x_12); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} } LEAN_EXPORT lean_object* l_Array_indexOfAux___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: @@ -1718,89 +1880,42 @@ return x_22; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_object* x_23; uint8_t x_24; lean_object* x_25; uint8_t x_104; lean_dec(x_18); x_23 = lean_ctor_get(x_17, 1); lean_inc(x_23); lean_dec(x_17); -x_24 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; -x_25 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_unbox(x_26); -lean_dec(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); -x_29 = lean_box(0); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_30 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_32 = !lean_is_exclusive(x_30); -if (x_32 == 0) +x_104 = l_Lean_Expr_isApp(x_15); +if (x_104 == 0) { -lean_object* x_33; -x_33 = lean_ctor_get(x_30, 0); -lean_dec(x_33); -return x_30; +x_24 = x_104; +x_25 = x_23; +goto block_103; } else { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_30, 1); -lean_inc(x_34); -lean_dec(x_30); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_31); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -else +lean_object* x_105; +lean_inc(x_15); +x_105 = l_Lean_Meta_Grind_isCongrRoot(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); +if (lean_obj_tag(x_105) == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_30, 1); -lean_inc(x_36); -lean_dec(x_30); -x_37 = lean_ctor_get(x_31, 0); -lean_inc(x_37); -lean_dec(x_31); -x_1 = x_16; -x_2 = x_37; -x_11 = x_36; -goto _start; -} +lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); +x_108 = lean_unbox(x_106); +lean_dec(x_106); +x_24 = x_108; +x_25 = x_107; +goto block_103; } else { -uint8_t x_39; +uint8_t x_109; lean_dec(x_16); +lean_dec(x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -1809,52 +1924,54 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_39 = !lean_is_exclusive(x_30); -if (x_39 == 0) +x_109 = !lean_is_exclusive(x_105); +if (x_109 == 0) { -return x_30; +return x_105; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_30, 0); -x_41 = lean_ctor_get(x_30, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_30); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_105, 0); +x_111 = lean_ctor_get(x_105, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_105); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } +block_103: +{ +if (x_24 == 0) +{ +lean_object* x_26; +lean_dec(x_15); +x_26 = lean_box(0); +x_1 = x_16; +x_2 = x_26; +x_11 = x_25; +goto _start; +} else { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_25); -if (x_43 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__4; +x_29 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_unbox(x_30); +lean_dec(x_30); +if (x_31 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_44 = lean_ctor_get(x_25, 1); -x_45 = lean_ctor_get(x_25, 0); -lean_dec(x_45); -lean_inc(x_15); -x_46 = l_Lean_MessageData_ofExpr(x_15); -x_47 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; -lean_ctor_set_tag(x_25, 7); -lean_ctor_set(x_25, 1, x_46); -lean_ctor_set(x_25, 0, x_47); -x_48 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_25); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_24, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_44); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -1863,16 +1980,15 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_53 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_51, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_52); -lean_dec(x_51); -if (lean_obj_tag(x_53) == 0) +x_34 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_54; -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -if (lean_obj_tag(x_54) == 0) +lean_object* x_35; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +if (lean_obj_tag(x_35) == 0) { -uint8_t x_55; +uint8_t x_36; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -1882,44 +1998,44 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_55 = !lean_is_exclusive(x_53); -if (x_55 == 0) +x_36 = !lean_is_exclusive(x_34); +if (x_36 == 0) { -lean_object* x_56; -x_56 = lean_ctor_get(x_53, 0); -lean_dec(x_56); -return x_53; +lean_object* x_37; +x_37 = lean_ctor_get(x_34, 0); +lean_dec(x_37); +return x_34; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_53, 1); -lean_inc(x_57); -lean_dec(x_53); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_54); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_34, 1); +lean_inc(x_38); +lean_dec(x_34); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_35); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } else { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_53, 1); -lean_inc(x_59); -lean_dec(x_53); -x_60 = lean_ctor_get(x_54, 0); -lean_inc(x_60); -lean_dec(x_54); +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_34, 1); +lean_inc(x_40); +lean_dec(x_34); +x_41 = lean_ctor_get(x_35, 0); +lean_inc(x_41); +lean_dec(x_35); x_1 = x_16; -x_2 = x_60; -x_11 = x_59; +x_2 = x_41; +x_11 = x_40; goto _start; } } else { -uint8_t x_62; +uint8_t x_43; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -1929,48 +2045,59 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_62 = !lean_is_exclusive(x_53); -if (x_62 == 0) +x_43 = !lean_is_exclusive(x_34); +if (x_43 == 0) { -return x_53; +return x_34; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_53, 0); -x_64 = lean_ctor_get(x_53, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_53); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_34, 0); +x_45 = lean_ctor_get(x_34, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_34); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_66 = lean_ctor_get(x_25, 1); -lean_inc(x_66); -lean_dec(x_25); +uint8_t x_47; +x_47 = !lean_is_exclusive(x_29); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_29, 1); +x_49 = lean_ctor_get(x_29, 0); +lean_dec(x_49); +x_50 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_48); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); lean_inc(x_15); -x_67 = l_Lean_MessageData_ofExpr(x_15); -x_68 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; -x_69 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_67); -x_70 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_71 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_24, x_71, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_66); -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); -lean_inc(x_74); -lean_dec(x_72); +x_52 = l_Lean_MessageData_ofExpr(x_15); +x_53 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; +lean_ctor_set_tag(x_29, 7); +lean_ctor_set(x_29, 1, x_52); +lean_ctor_set(x_29, 0, x_53); +x_54 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_29); +lean_ctor_set(x_55, 1, x_54); +x_56 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_28, x_55, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_51); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -1979,16 +2106,16 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_75 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_73, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_74); -lean_dec(x_73); -if (lean_obj_tag(x_75) == 0) +x_59 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_57, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_58); +lean_dec(x_57); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_76; -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -if (lean_obj_tag(x_76) == 0) +lean_object* x_60; +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; +uint8_t x_61; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -1998,43 +2125,44 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_78 = x_75; -} else { - lean_dec_ref(x_75); - x_78 = lean_box(0); +x_61 = !lean_is_exclusive(x_59); +if (x_61 == 0) +{ +lean_object* x_62; +x_62 = lean_ctor_get(x_59, 0); +lean_dec(x_62); +return x_59; } -if (lean_is_scalar(x_78)) { - x_79 = lean_alloc_ctor(0, 2, 0); -} else { - x_79 = x_78; +else +{ +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_59, 1); +lean_inc(x_63); +lean_dec(x_59); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_60); +lean_ctor_set(x_64, 1, x_63); +return x_64; } -lean_ctor_set(x_79, 0, x_76); -lean_ctor_set(x_79, 1, x_77); -return x_79; } else { -lean_object* x_80; lean_object* x_81; -x_80 = lean_ctor_get(x_75, 1); -lean_inc(x_80); -lean_dec(x_75); -x_81 = lean_ctor_get(x_76, 0); -lean_inc(x_81); -lean_dec(x_76); +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_59, 1); +lean_inc(x_65); +lean_dec(x_59); +x_66 = lean_ctor_get(x_60, 0); +lean_inc(x_66); +lean_dec(x_60); x_1 = x_16; -x_2 = x_81; -x_11 = x_80; +x_2 = x_66; +x_11 = x_65; goto _start; } } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +uint8_t x_68; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -2044,26 +2172,219 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_83 = lean_ctor_get(x_75, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_75, 1); -lean_inc(x_84); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_85 = x_75; -} else { - lean_dec_ref(x_75); - x_85 = lean_box(0); +x_68 = !lean_is_exclusive(x_59); +if (x_68 == 0) +{ +return x_59; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_59, 0); +x_70 = lean_ctor_get(x_59, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_59); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +else +{ +uint8_t x_72; +lean_free_object(x_29); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_72 = !lean_is_exclusive(x_50); +if (x_72 == 0) +{ +return x_50; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_50, 0); +x_74 = lean_ctor_get(x_50, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_50); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; +} +} +} +else +{ +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_29, 1); +lean_inc(x_76); +lean_dec(x_29); +x_77 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_76); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +lean_inc(x_15); +x_79 = l_Lean_MessageData_ofExpr(x_15); +x_80 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___closed__2; +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +x_82 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_28, x_83, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_78); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_87 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_reinsertParents___spec__1___lambda__1(x_15, x_85, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_86); +lean_dec(x_85); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; +} else { + lean_dec_ref(x_87); + x_90 = lean_box(0); +} +if (lean_is_scalar(x_90)) { + x_91 = lean_alloc_ctor(0, 2, 0); +} else { + x_91 = x_90; +} +lean_ctor_set(x_91, 0, x_88); +lean_ctor_set(x_91, 1, x_89); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_87, 1); +lean_inc(x_92); +lean_dec(x_87); +x_93 = lean_ctor_get(x_88, 0); +lean_inc(x_93); +lean_dec(x_88); +x_1 = x_16; +x_2 = x_93; +x_11 = x_92; +goto _start; +} +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_95 = lean_ctor_get(x_87, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_87, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_97 = x_87; +} else { + lean_dec_ref(x_87); + x_97 = lean_box(0); +} +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(1, 2, 0); +} else { + x_98 = x_97; +} +lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 1, x_96); +return x_98; +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_99 = lean_ctor_get(x_77, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_77, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_101 = x_77; +} else { + lean_dec_ref(x_77); + x_101 = lean_box(0); +} +if (lean_is_scalar(x_101)) { + x_102 = lean_alloc_ctor(1, 2, 0); +} else { + x_102 = x_101; +} +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_100); +return x_102; } -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(1, 2, 0); -} else { - x_86 = x_85; } -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_84); -return x_86; } } } @@ -2071,7 +2392,7 @@ return x_86; } else { -uint8_t x_87; +uint8_t x_113; lean_dec(x_16); lean_dec(x_15); lean_dec(x_10); @@ -2082,23 +2403,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_87 = !lean_is_exclusive(x_17); -if (x_87 == 0) +x_113 = !lean_is_exclusive(x_17); +if (x_113 == 0) { return x_17; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_17, 0); -x_89 = lean_ctor_get(x_17, 1); -lean_inc(x_89); -lean_inc(x_88); +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_17, 0); +x_115 = lean_ctor_get(x_17, 1); +lean_inc(x_115); +lean_inc(x_114); lean_dec(x_17); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_88); -lean_ctor_set(x_90, 1, x_89); -return x_90; +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; } } } @@ -2266,10 +2587,6 @@ x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); x_28 = l_Lean_Meta_Grind_closeGoal(x_26, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); return x_28; } else @@ -2591,10 +2908,6 @@ x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); x_28 = l_Lean_Meta_Grind_closeGoal(x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); return x_28; } else @@ -4363,7 +4676,7 @@ x_29 = l_Lean_Meta_Grind_setENode(x_1, x_2, x_12, x_13, x_14, x_15, x_16, x_17, x_30 = !lean_is_exclusive(x_29); if (x_30 == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; uint8_t x_39; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; x_31 = lean_ctor_get(x_29, 1); x_32 = lean_ctor_get(x_29, 0); lean_dec(x_32); @@ -4377,14 +4690,16 @@ x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); lean_dec(x_6); x_38 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_33, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_31); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +if (lean_obj_tag(x_38) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_ctor_get(x_38, 0); -x_41 = lean_ctor_get(x_38, 1); -x_42 = lean_ctor_get(x_7, 2); -lean_inc(x_42); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_7, 2); +lean_inc(x_41); lean_dec(x_7); lean_inc(x_19); lean_inc(x_18); @@ -4394,167 +4709,172 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_42); +lean_inc(x_41); lean_inc(x_1); -x_43 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_42, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_41); -if (lean_obj_tag(x_43) == 0) +x_42 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_41, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_40); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); -lean_dec(x_43); +lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); lean_inc(x_8); -x_45 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_44); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_unbox(x_46); -lean_dec(x_46); -if (x_47 == 0) +x_44 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_43); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_unbox(x_45); +lean_dec(x_45); +if (x_46 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -lean_free_object(x_38); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_free_object(x_29); lean_dec(x_8); lean_dec(x_1); -x_48 = lean_ctor_get(x_45, 1); -lean_inc(x_48); -lean_dec(x_45); -x_49 = lean_box(0); -x_50 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_49, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_48); +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_box(0); +x_49 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_48, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_47); lean_dec(x_35); -return x_50; +return x_49; } else { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_45); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_52 = lean_ctor_get(x_45, 1); -x_53 = lean_ctor_get(x_45, 0); -lean_dec(x_53); -x_54 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_52); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) +uint8_t x_50; +x_50 = !lean_is_exclusive(x_44); +if (x_50 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_56 = lean_ctor_get(x_54, 0); -x_57 = lean_ctor_get(x_54, 1); -x_58 = l_Lean_Meta_Grind_ppENodeRef(x_42, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_57); -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_44, 1); +x_52 = lean_ctor_get(x_44, 0); +lean_dec(x_52); +x_53 = l_Lean_Meta_Grind_updateLastTag(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_51); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_58, 0); -x_61 = lean_ctor_get(x_58, 1); -x_62 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_61); -if (lean_obj_tag(x_62) == 0) +lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_54); +x_56 = !lean_is_exclusive(x_55); +if (x_56 == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); +lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_57 = lean_ctor_get(x_55, 0); +x_58 = lean_ctor_get(x_55, 1); +x_59 = l_Lean_Meta_Grind_ppENodeRef(x_41, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_58); +x_60 = !lean_is_exclusive(x_59); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_59, 0); +x_62 = lean_ctor_get(x_59, 1); +x_63 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_62); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); -lean_dec(x_62); -x_65 = l_Lean_Meta_Grind_ppENodeRef(x_63, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); lean_dec(x_63); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) +x_66 = l_Lean_Meta_Grind_ppENodeRef(x_64, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_65); +lean_dec(x_64); +x_67 = !lean_is_exclusive(x_66); +if (x_67 == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_67 = lean_ctor_get(x_65, 0); -x_68 = lean_ctor_get(x_65, 1); -x_69 = l_Lean_MessageData_ofFormat(x_56); -x_70 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_65, 7); -lean_ctor_set(x_65, 1, x_69); -lean_ctor_set(x_65, 0, x_70); -x_71 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -lean_ctor_set_tag(x_58, 7); -lean_ctor_set(x_58, 1, x_71); -lean_ctor_set(x_58, 0, x_65); -x_72 = l_Lean_MessageData_ofFormat(x_60); -lean_ctor_set_tag(x_54, 7); -lean_ctor_set(x_54, 1, x_72); -lean_ctor_set(x_54, 0, x_58); -x_73 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -lean_ctor_set_tag(x_45, 7); -lean_ctor_set(x_45, 1, x_73); -lean_ctor_set(x_45, 0, x_54); -x_74 = l_Lean_MessageData_ofFormat(x_67); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_74); -lean_ctor_set(x_38, 0, x_45); +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_68 = lean_ctor_get(x_66, 0); +x_69 = lean_ctor_get(x_66, 1); +x_70 = l_Lean_MessageData_ofFormat(x_57); +x_71 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_66, 7); +lean_ctor_set(x_66, 1, x_70); +lean_ctor_set(x_66, 0, x_71); +x_72 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +lean_ctor_set_tag(x_59, 7); +lean_ctor_set(x_59, 1, x_72); +lean_ctor_set(x_59, 0, x_66); +x_73 = l_Lean_MessageData_ofFormat(x_61); +lean_ctor_set_tag(x_55, 7); +lean_ctor_set(x_55, 1, x_73); +lean_ctor_set(x_55, 0, x_59); +x_74 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_74); +lean_ctor_set(x_44, 0, x_55); +x_75 = l_Lean_MessageData_ofFormat(x_68); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_70); -lean_ctor_set(x_29, 0, x_38); -x_75 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_68); -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -lean_dec(x_75); -x_78 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_76, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_77); -lean_dec(x_76); +lean_ctor_set(x_29, 1, x_75); +lean_ctor_set(x_29, 0, x_44); +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_29); +lean_ctor_set(x_76, 1, x_71); +x_77 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_76, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_69); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +lean_dec(x_77); +x_80 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_78, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_79); +lean_dec(x_78); lean_dec(x_35); -return x_78; +return x_80; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_79 = lean_ctor_get(x_65, 0); -x_80 = lean_ctor_get(x_65, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_65); -x_81 = l_Lean_MessageData_ofFormat(x_56); -x_82 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_83 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_81); -x_84 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -lean_ctor_set_tag(x_58, 7); -lean_ctor_set(x_58, 1, x_84); -lean_ctor_set(x_58, 0, x_83); -x_85 = l_Lean_MessageData_ofFormat(x_60); -lean_ctor_set_tag(x_54, 7); -lean_ctor_set(x_54, 1, x_85); -lean_ctor_set(x_54, 0, x_58); -x_86 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -lean_ctor_set_tag(x_45, 7); -lean_ctor_set(x_45, 1, x_86); -lean_ctor_set(x_45, 0, x_54); -x_87 = l_Lean_MessageData_ofFormat(x_79); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_87); -lean_ctor_set(x_38, 0, x_45); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_81 = lean_ctor_get(x_66, 0); +x_82 = lean_ctor_get(x_66, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_66); +x_83 = l_Lean_MessageData_ofFormat(x_57); +x_84 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_83); +x_86 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +lean_ctor_set_tag(x_59, 7); +lean_ctor_set(x_59, 1, x_86); +lean_ctor_set(x_59, 0, x_85); +x_87 = l_Lean_MessageData_ofFormat(x_61); +lean_ctor_set_tag(x_55, 7); +lean_ctor_set(x_55, 1, x_87); +lean_ctor_set(x_55, 0, x_59); +x_88 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_88); +lean_ctor_set(x_44, 0, x_55); +x_89 = l_Lean_MessageData_ofFormat(x_81); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_82); -lean_ctor_set(x_29, 0, x_38); -x_88 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_80); -x_89 = lean_ctor_get(x_88, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_88, 1); -lean_inc(x_90); -lean_dec(x_88); -x_91 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_89, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_90); -lean_dec(x_89); +lean_ctor_set(x_29, 1, x_89); +lean_ctor_set(x_29, 0, x_44); +x_90 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_90, 0, x_29); +lean_ctor_set(x_90, 1, x_84); +x_91 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_90, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_82); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +lean_dec(x_91); +x_94 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_92, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_93); +lean_dec(x_92); lean_dec(x_35); -return x_91; +return x_94; } } else { -uint8_t x_92; -lean_free_object(x_58); -lean_dec(x_60); -lean_free_object(x_54); -lean_dec(x_56); -lean_free_object(x_45); -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +uint8_t x_95; +lean_free_object(x_59); +lean_dec(x_61); +lean_free_object(x_55); +lean_dec(x_57); +lean_free_object(x_44); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -4570,107 +4890,106 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_92 = !lean_is_exclusive(x_62); -if (x_92 == 0) +x_95 = !lean_is_exclusive(x_63); +if (x_95 == 0) { -return x_62; +return x_63; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_62, 0); -x_94 = lean_ctor_get(x_62, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_62); -x_95 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_94); -return x_95; +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_63, 0); +x_97 = lean_ctor_get(x_63, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_63); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; } } } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_58, 0); -x_97 = lean_ctor_get(x_58, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_58); -x_98 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_97); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_59, 0); +x_100 = lean_ctor_get(x_59, 1); lean_inc(x_100); -lean_dec(x_98); -x_101 = l_Lean_Meta_Grind_ppENodeRef(x_99, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_100); -lean_dec(x_99); +lean_inc(x_99); +lean_dec(x_59); +x_101 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_100); +if (lean_obj_tag(x_101) == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; x_102 = lean_ctor_get(x_101, 0); lean_inc(x_102); x_103 = lean_ctor_get(x_101, 1); lean_inc(x_103); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_104 = x_101; +lean_dec(x_101); +x_104 = l_Lean_Meta_Grind_ppENodeRef(x_102, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_103); +lean_dec(x_102); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_107 = x_104; } else { - lean_dec_ref(x_101); - x_104 = lean_box(0); + lean_dec_ref(x_104); + x_107 = lean_box(0); } -x_105 = l_Lean_MessageData_ofFormat(x_56); -x_106 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_104)) { - x_107 = lean_alloc_ctor(7, 2, 0); +x_108 = l_Lean_MessageData_ofFormat(x_57); +x_109 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_107)) { + x_110 = lean_alloc_ctor(7, 2, 0); } else { - x_107 = x_104; - lean_ctor_set_tag(x_107, 7); -} -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_105); -x_108 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -x_109 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_MessageData_ofFormat(x_96); -lean_ctor_set_tag(x_54, 7); -lean_ctor_set(x_54, 1, x_110); -lean_ctor_set(x_54, 0, x_109); -x_111 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -lean_ctor_set_tag(x_45, 7); -lean_ctor_set(x_45, 1, x_111); -lean_ctor_set(x_45, 0, x_54); -x_112 = l_Lean_MessageData_ofFormat(x_102); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_112); -lean_ctor_set(x_38, 0, x_45); + x_110 = x_107; + lean_ctor_set_tag(x_110, 7); +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_108); +x_111 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +x_112 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +x_113 = l_Lean_MessageData_ofFormat(x_99); +lean_ctor_set_tag(x_55, 7); +lean_ctor_set(x_55, 1, x_113); +lean_ctor_set(x_55, 0, x_112); +x_114 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_114); +lean_ctor_set(x_44, 0, x_55); +x_115 = l_Lean_MessageData_ofFormat(x_105); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_106); -lean_ctor_set(x_29, 0, x_38); -x_113 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_103); -x_114 = lean_ctor_get(x_113, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -lean_dec(x_113); -x_116 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_114, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_115); -lean_dec(x_114); +lean_ctor_set(x_29, 1, x_115); +lean_ctor_set(x_29, 0, x_44); +x_116 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_116, 0, x_29); +lean_ctor_set(x_116, 1, x_109); +x_117 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_116, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_106); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec(x_117); +x_120 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_118, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_119); +lean_dec(x_118); lean_dec(x_35); -return x_116; +return x_120; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -lean_dec(x_96); -lean_free_object(x_54); -lean_dec(x_56); -lean_free_object(x_45); -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_99); +lean_free_object(x_55); +lean_dec(x_57); +lean_free_object(x_44); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -4686,128 +5005,127 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_117 = lean_ctor_get(x_98, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_98, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_119 = x_98; +x_121 = lean_ctor_get(x_101, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_101, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_123 = x_101; } else { - lean_dec_ref(x_98); - x_119 = lean_box(0); + lean_dec_ref(x_101); + x_123 = lean_box(0); } -if (lean_is_scalar(x_119)) { - x_120 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_123)) { + x_124 = lean_alloc_ctor(1, 2, 0); } else { - x_120 = x_119; + x_124 = x_123; } -lean_ctor_set(x_120, 0, x_117); -lean_ctor_set(x_120, 1, x_118); -return x_120; +lean_ctor_set(x_124, 0, x_121); +lean_ctor_set(x_124, 1, x_122); +return x_124; } } } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_121 = lean_ctor_get(x_54, 0); -x_122 = lean_ctor_get(x_54, 1); -lean_inc(x_122); -lean_inc(x_121); -lean_dec(x_54); -x_123 = l_Lean_Meta_Grind_ppENodeRef(x_42, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_122); -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_123, 1); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_125 = lean_ctor_get(x_55, 0); +x_126 = lean_ctor_get(x_55, 1); +lean_inc(x_126); lean_inc(x_125); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_126 = x_123; -} else { - lean_dec_ref(x_123); - x_126 = lean_box(0); -} -x_127 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_125); -if (lean_obj_tag(x_127) == 0) -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_55); +x_127 = l_Lean_Meta_Grind_ppENodeRef(x_41, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_126); x_128 = lean_ctor_get(x_127, 0); lean_inc(x_128); x_129 = lean_ctor_get(x_127, 1); lean_inc(x_129); -lean_dec(x_127); -x_130 = l_Lean_Meta_Grind_ppENodeRef(x_128, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_129); -lean_dec(x_128); -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_130, 1); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_130 = x_127; +} else { + lean_dec_ref(x_127); + x_130 = lean_box(0); +} +x_131 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_129); +if (lean_obj_tag(x_131) == 0) +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_132 = lean_ctor_get(x_131, 0); lean_inc(x_132); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_133 = x_130; +x_133 = lean_ctor_get(x_131, 1); +lean_inc(x_133); +lean_dec(x_131); +x_134 = l_Lean_Meta_Grind_ppENodeRef(x_132, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_133); +lean_dec(x_132); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_137 = x_134; } else { - lean_dec_ref(x_130); - x_133 = lean_box(0); + lean_dec_ref(x_134); + x_137 = lean_box(0); } -x_134 = l_Lean_MessageData_ofFormat(x_121); -x_135 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_133)) { - x_136 = lean_alloc_ctor(7, 2, 0); +x_138 = l_Lean_MessageData_ofFormat(x_125); +x_139 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_137)) { + x_140 = lean_alloc_ctor(7, 2, 0); } else { - x_136 = x_133; - lean_ctor_set_tag(x_136, 7); -} -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_134); -x_137 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_126)) { - x_138 = lean_alloc_ctor(7, 2, 0); + x_140 = x_137; + lean_ctor_set_tag(x_140, 7); +} +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_138); +x_141 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +if (lean_is_scalar(x_130)) { + x_142 = lean_alloc_ctor(7, 2, 0); } else { - x_138 = x_126; - lean_ctor_set_tag(x_138, 7); -} -lean_ctor_set(x_138, 0, x_136); -lean_ctor_set(x_138, 1, x_137); -x_139 = l_Lean_MessageData_ofFormat(x_124); -x_140 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_140, 0, x_138); -lean_ctor_set(x_140, 1, x_139); -x_141 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -lean_ctor_set_tag(x_45, 7); -lean_ctor_set(x_45, 1, x_141); -lean_ctor_set(x_45, 0, x_140); -x_142 = l_Lean_MessageData_ofFormat(x_131); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_142); -lean_ctor_set(x_38, 0, x_45); + x_142 = x_130; + lean_ctor_set_tag(x_142, 7); +} +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Lean_MessageData_ofFormat(x_128); +x_144 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +x_145 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +lean_ctor_set_tag(x_44, 7); +lean_ctor_set(x_44, 1, x_145); +lean_ctor_set(x_44, 0, x_144); +x_146 = l_Lean_MessageData_ofFormat(x_135); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_135); -lean_ctor_set(x_29, 0, x_38); -x_143 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_132); -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_143, 1); -lean_inc(x_145); -lean_dec(x_143); -x_146 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_144, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_145); -lean_dec(x_144); +lean_ctor_set(x_29, 1, x_146); +lean_ctor_set(x_29, 0, x_44); +x_147 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_147, 0, x_29); +lean_ctor_set(x_147, 1, x_139); +x_148 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_147, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_136); +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +lean_dec(x_148); +x_151 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_149, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_150); +lean_dec(x_149); lean_dec(x_35); -return x_146; +return x_151; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -lean_dec(x_126); -lean_dec(x_124); -lean_dec(x_121); -lean_free_object(x_45); -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +lean_dec(x_130); +lean_dec(x_128); +lean_dec(x_125); +lean_free_object(x_44); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -4823,72 +5141,85 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_147 = lean_ctor_get(x_127, 0); -lean_inc(x_147); -x_148 = lean_ctor_get(x_127, 1); -lean_inc(x_148); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - lean_ctor_release(x_127, 1); - x_149 = x_127; +x_152 = lean_ctor_get(x_131, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_131, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_154 = x_131; } else { - lean_dec_ref(x_127); - x_149 = lean_box(0); + lean_dec_ref(x_131); + x_154 = lean_box(0); } -if (lean_is_scalar(x_149)) { - x_150 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_154)) { + x_155 = lean_alloc_ctor(1, 2, 0); } else { - x_150 = x_149; + x_155 = x_154; } -lean_ctor_set(x_150, 0, x_147); -lean_ctor_set(x_150, 1, x_148); -return x_150; +lean_ctor_set(x_155, 0, x_152); +lean_ctor_set(x_155, 1, x_153); +return x_155; } } } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_151 = lean_ctor_get(x_45, 1); -lean_inc(x_151); -lean_dec(x_45); -x_152 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_151); -x_153 = lean_ctor_get(x_152, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_152, 1); -lean_inc(x_154); -if (lean_is_exclusive(x_152)) { - lean_ctor_release(x_152, 0); - lean_ctor_release(x_152, 1); - x_155 = x_152; -} else { - lean_dec_ref(x_152); - x_155 = lean_box(0); +uint8_t x_156; +lean_free_object(x_44); +lean_dec(x_41); +lean_dec(x_39); +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_33); +lean_free_object(x_29); +lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_156 = !lean_is_exclusive(x_53); +if (x_156 == 0) +{ +return x_53; } -x_156 = l_Lean_Meta_Grind_ppENodeRef(x_42, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_154); -x_157 = lean_ctor_get(x_156, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_156, 1); +else +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_157 = lean_ctor_get(x_53, 0); +x_158 = lean_ctor_get(x_53, 1); lean_inc(x_158); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - lean_ctor_release(x_156, 1); - x_159 = x_156; -} else { - lean_dec_ref(x_156); - x_159 = lean_box(0); +lean_inc(x_157); +lean_dec(x_53); +x_159 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +return x_159; } -x_160 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_158); -if (lean_obj_tag(x_160) == 0) +} +} +else +{ +lean_object* x_160; lean_object* x_161; +x_160 = lean_ctor_get(x_44, 1); +lean_inc(x_160); +lean_dec(x_44); +x_161 = l_Lean_Meta_Grind_updateLastTag(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_160); +if (lean_obj_tag(x_161) == 0) { -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_160, 1); +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_162 = lean_ctor_get(x_161, 1); lean_inc(x_162); -lean_dec(x_160); -x_163 = l_Lean_Meta_Grind_ppENodeRef(x_161, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_162); lean_dec(x_161); +x_163 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_162); x_164 = lean_ctor_get(x_163, 0); lean_inc(x_164); x_165 = lean_ctor_get(x_163, 1); @@ -4901,66 +5232,101 @@ if (lean_is_exclusive(x_163)) { lean_dec_ref(x_163); x_166 = lean_box(0); } -x_167 = l_Lean_MessageData_ofFormat(x_153); -x_168 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_166)) { - x_169 = lean_alloc_ctor(7, 2, 0); +x_167 = l_Lean_Meta_Grind_ppENodeRef(x_41, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_165); +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_167, 1); +lean_inc(x_169); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_170 = x_167; +} else { + lean_dec_ref(x_167); + x_170 = lean_box(0); +} +x_171 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_169); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = l_Lean_Meta_Grind_ppENodeRef(x_172, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_173); +lean_dec(x_172); +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_174, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_177 = x_174; +} else { + lean_dec_ref(x_174); + x_177 = lean_box(0); +} +x_178 = l_Lean_MessageData_ofFormat(x_164); +x_179 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_177)) { + x_180 = lean_alloc_ctor(7, 2, 0); } else { - x_169 = x_166; - lean_ctor_set_tag(x_169, 7); -} -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_167); -x_170 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_159)) { - x_171 = lean_alloc_ctor(7, 2, 0); + x_180 = x_177; + lean_ctor_set_tag(x_180, 7); +} +lean_ctor_set(x_180, 0, x_179); +lean_ctor_set(x_180, 1, x_178); +x_181 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +if (lean_is_scalar(x_170)) { + x_182 = lean_alloc_ctor(7, 2, 0); } else { - x_171 = x_159; - lean_ctor_set_tag(x_171, 7); -} -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -x_172 = l_Lean_MessageData_ofFormat(x_157); -if (lean_is_scalar(x_155)) { - x_173 = lean_alloc_ctor(7, 2, 0); + x_182 = x_170; + lean_ctor_set_tag(x_182, 7); +} +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +x_183 = l_Lean_MessageData_ofFormat(x_168); +if (lean_is_scalar(x_166)) { + x_184 = lean_alloc_ctor(7, 2, 0); } else { - x_173 = x_155; - lean_ctor_set_tag(x_173, 7); -} -lean_ctor_set(x_173, 0, x_171); -lean_ctor_set(x_173, 1, x_172); -x_174 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -x_175 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_175, 0, x_173); -lean_ctor_set(x_175, 1, x_174); -x_176 = l_Lean_MessageData_ofFormat(x_164); -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_176); -lean_ctor_set(x_38, 0, x_175); + x_184 = x_166; + lean_ctor_set_tag(x_184, 7); +} +lean_ctor_set(x_184, 0, x_182); +lean_ctor_set(x_184, 1, x_183); +x_185 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +x_186 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +x_187 = l_Lean_MessageData_ofFormat(x_175); lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_168); -lean_ctor_set(x_29, 0, x_38); -x_177 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_165); -x_178 = lean_ctor_get(x_177, 0); -lean_inc(x_178); -x_179 = lean_ctor_get(x_177, 1); -lean_inc(x_179); -lean_dec(x_177); -x_180 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_40, x_33, x_9, x_24, x_35, x_34, x_42, x_10, x_37, x_36, x_178, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_179); -lean_dec(x_178); +lean_ctor_set(x_29, 1, x_187); +lean_ctor_set(x_29, 0, x_186); +x_188 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_188, 0, x_29); +lean_ctor_set(x_188, 1, x_179); +x_189 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_188, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_176); +x_190 = lean_ctor_get(x_189, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_189, 1); +lean_inc(x_191); +lean_dec(x_189); +x_192 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_39, x_33, x_9, x_24, x_35, x_34, x_41, x_10, x_37, x_36, x_190, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_191); +lean_dec(x_190); lean_dec(x_35); -return x_180; +return x_192; } else { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -lean_dec(x_159); -lean_dec(x_157); -lean_dec(x_155); -lean_dec(x_153); -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; +lean_dec(x_170); +lean_dec(x_168); +lean_dec(x_166); +lean_dec(x_164); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -4976,36 +5342,33 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_181 = lean_ctor_get(x_160, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_160, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_160)) { - lean_ctor_release(x_160, 0); - lean_ctor_release(x_160, 1); - x_183 = x_160; +x_193 = lean_ctor_get(x_171, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_171, 1); +lean_inc(x_194); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_195 = x_171; } else { - lean_dec_ref(x_160); - x_183 = lean_box(0); + lean_dec_ref(x_171); + x_195 = lean_box(0); } -if (lean_is_scalar(x_183)) { - x_184 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_195)) { + x_196 = lean_alloc_ctor(1, 2, 0); } else { - x_184 = x_183; -} -lean_ctor_set(x_184, 0, x_181); -lean_ctor_set(x_184, 1, x_182); -return x_184; -} + x_196 = x_195; } +lean_ctor_set(x_196, 0, x_193); +lean_ctor_set(x_196, 1, x_194); +return x_196; } } else { -uint8_t x_185; -lean_dec(x_42); -lean_free_object(x_38); -lean_dec(x_40); +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -5022,201 +5385,35 @@ lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_1); -x_185 = !lean_is_exclusive(x_43); -if (x_185 == 0) -{ -return x_43; -} -else -{ -lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_186 = lean_ctor_get(x_43, 0); -x_187 = lean_ctor_get(x_43, 1); -lean_inc(x_187); -lean_inc(x_186); -lean_dec(x_43); -x_188 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_188, 0, x_186); -lean_ctor_set(x_188, 1, x_187); -return x_188; -} -} -} -else -{ -lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_189 = lean_ctor_get(x_38, 0); -x_190 = lean_ctor_get(x_38, 1); -lean_inc(x_190); -lean_inc(x_189); -lean_dec(x_38); -x_191 = lean_ctor_get(x_7, 2); -lean_inc(x_191); -lean_dec(x_7); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_191); -lean_inc(x_1); -x_192 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_191, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_190); -if (lean_obj_tag(x_192) == 0) -{ -lean_object* x_193; lean_object* x_194; lean_object* x_195; uint8_t x_196; -x_193 = lean_ctor_get(x_192, 1); -lean_inc(x_193); -lean_dec(x_192); -lean_inc(x_8); -x_194 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_193); -x_195 = lean_ctor_get(x_194, 0); -lean_inc(x_195); -x_196 = lean_unbox(x_195); -lean_dec(x_195); -if (x_196 == 0) -{ -lean_object* x_197; lean_object* x_198; lean_object* x_199; -lean_free_object(x_29); -lean_dec(x_8); -lean_dec(x_1); -x_197 = lean_ctor_get(x_194, 1); +x_197 = lean_ctor_get(x_161, 0); lean_inc(x_197); -lean_dec(x_194); -x_198 = lean_box(0); -x_199 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_189, x_33, x_9, x_24, x_35, x_34, x_191, x_10, x_37, x_36, x_198, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_197); -lean_dec(x_35); -return x_199; -} -else -{ -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; -x_200 = lean_ctor_get(x_194, 1); -lean_inc(x_200); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - x_201 = x_194; +x_198 = lean_ctor_get(x_161, 1); +lean_inc(x_198); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_199 = x_161; } else { - lean_dec_ref(x_194); - x_201 = lean_box(0); + lean_dec_ref(x_161); + x_199 = lean_box(0); } -x_202 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_200); -x_203 = lean_ctor_get(x_202, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_202, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - lean_ctor_release(x_202, 1); - x_205 = x_202; +if (lean_is_scalar(x_199)) { + x_200 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_202); - x_205 = lean_box(0); + x_200 = x_199; +} +lean_ctor_set(x_200, 0, x_197); +lean_ctor_set(x_200, 1, x_198); +return x_200; } -x_206 = l_Lean_Meta_Grind_ppENodeRef(x_191, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_204); -x_207 = lean_ctor_get(x_206, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_206, 1); -lean_inc(x_208); -if (lean_is_exclusive(x_206)) { - lean_ctor_release(x_206, 0); - lean_ctor_release(x_206, 1); - x_209 = x_206; -} else { - lean_dec_ref(x_206); - x_209 = lean_box(0); } -x_210 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_208); -if (lean_obj_tag(x_210) == 0) -{ -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_210, 1); -lean_inc(x_212); -lean_dec(x_210); -x_213 = l_Lean_Meta_Grind_ppENodeRef(x_211, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_212); -lean_dec(x_211); -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_216 = x_213; -} else { - lean_dec_ref(x_213); - x_216 = lean_box(0); } -x_217 = l_Lean_MessageData_ofFormat(x_203); -x_218 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_216)) { - x_219 = lean_alloc_ctor(7, 2, 0); -} else { - x_219 = x_216; - lean_ctor_set_tag(x_219, 7); -} -lean_ctor_set(x_219, 0, x_218); -lean_ctor_set(x_219, 1, x_217); -x_220 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_209)) { - x_221 = lean_alloc_ctor(7, 2, 0); -} else { - x_221 = x_209; - lean_ctor_set_tag(x_221, 7); -} -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -x_222 = l_Lean_MessageData_ofFormat(x_207); -if (lean_is_scalar(x_205)) { - x_223 = lean_alloc_ctor(7, 2, 0); -} else { - x_223 = x_205; - lean_ctor_set_tag(x_223, 7); -} -lean_ctor_set(x_223, 0, x_221); -lean_ctor_set(x_223, 1, x_222); -x_224 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -if (lean_is_scalar(x_201)) { - x_225 = lean_alloc_ctor(7, 2, 0); -} else { - x_225 = x_201; - lean_ctor_set_tag(x_225, 7); -} -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -x_226 = l_Lean_MessageData_ofFormat(x_214); -x_227 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_227, 0, x_225); -lean_ctor_set(x_227, 1, x_226); -lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_218); -lean_ctor_set(x_29, 0, x_227); -x_228 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_29, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_215); -x_229 = lean_ctor_get(x_228, 0); -lean_inc(x_229); -x_230 = lean_ctor_get(x_228, 1); -lean_inc(x_230); -lean_dec(x_228); -x_231 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_189, x_33, x_9, x_24, x_35, x_34, x_191, x_10, x_37, x_36, x_229, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_230); -lean_dec(x_229); -lean_dec(x_35); -return x_231; } else { -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; -lean_dec(x_209); -lean_dec(x_207); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_201); -lean_dec(x_191); -lean_dec(x_189); +uint8_t x_201; +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -5232,34 +5429,30 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_232 = lean_ctor_get(x_210, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_210, 1); -lean_inc(x_233); -if (lean_is_exclusive(x_210)) { - lean_ctor_release(x_210, 0); - lean_ctor_release(x_210, 1); - x_234 = x_210; -} else { - lean_dec_ref(x_210); - x_234 = lean_box(0); -} -if (lean_is_scalar(x_234)) { - x_235 = lean_alloc_ctor(1, 2, 0); -} else { - x_235 = x_234; +lean_dec(x_1); +x_201 = !lean_is_exclusive(x_42); +if (x_201 == 0) +{ +return x_42; } -lean_ctor_set(x_235, 0, x_232); -lean_ctor_set(x_235, 1, x_233); -return x_235; +else +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_202 = lean_ctor_get(x_42, 0); +x_203 = lean_ctor_get(x_42, 1); +lean_inc(x_203); +lean_inc(x_202); +lean_dec(x_42); +x_204 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_204, 0, x_202); +lean_ctor_set(x_204, 1, x_203); +return x_204; } } } else { -lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; -lean_dec(x_191); -lean_dec(x_189); +uint8_t x_205; lean_dec(x_35); lean_dec(x_34); lean_dec(x_33); @@ -5275,60 +5468,54 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_1); -x_236 = lean_ctor_get(x_192, 0); -lean_inc(x_236); -x_237 = lean_ctor_get(x_192, 1); -lean_inc(x_237); -if (lean_is_exclusive(x_192)) { - lean_ctor_release(x_192, 0); - lean_ctor_release(x_192, 1); - x_238 = x_192; -} else { - lean_dec_ref(x_192); - x_238 = lean_box(0); -} -if (lean_is_scalar(x_238)) { - x_239 = lean_alloc_ctor(1, 2, 0); -} else { - x_239 = x_238; +x_205 = !lean_is_exclusive(x_38); +if (x_205 == 0) +{ +return x_38; } -lean_ctor_set(x_239, 0, x_236); -lean_ctor_set(x_239, 1, x_237); -return x_239; +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; +x_206 = lean_ctor_get(x_38, 0); +x_207 = lean_ctor_get(x_38, 1); +lean_inc(x_207); +lean_inc(x_206); +lean_dec(x_38); +x_208 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_208, 0, x_206); +lean_ctor_set(x_208, 1, x_207); +return x_208; } } } else { -lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; uint8_t x_244; uint8_t x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -x_240 = lean_ctor_get(x_29, 1); -lean_inc(x_240); +lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; uint8_t x_214; lean_object* x_215; +x_209 = lean_ctor_get(x_29, 1); +lean_inc(x_209); lean_dec(x_29); -x_241 = lean_ctor_get(x_6, 0); -lean_inc(x_241); -x_242 = lean_ctor_get(x_6, 1); -lean_inc(x_242); -x_243 = lean_ctor_get(x_6, 6); -lean_inc(x_243); -x_244 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); -x_245 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); +x_210 = lean_ctor_get(x_6, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_6, 1); +lean_inc(x_211); +x_212 = lean_ctor_get(x_6, 6); +lean_inc(x_212); +x_213 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); +x_214 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); lean_dec(x_6); -x_246 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_241, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_240); -x_247 = lean_ctor_get(x_246, 0); -lean_inc(x_247); -x_248 = lean_ctor_get(x_246, 1); -lean_inc(x_248); -if (lean_is_exclusive(x_246)) { - lean_ctor_release(x_246, 0); - lean_ctor_release(x_246, 1); - x_249 = x_246; -} else { - lean_dec_ref(x_246); - x_249 = lean_box(0); -} -x_250 = lean_ctor_get(x_7, 2); -lean_inc(x_250); +x_215 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_210, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_209); +if (lean_obj_tag(x_215) == 0) +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; +x_216 = lean_ctor_get(x_215, 0); +lean_inc(x_216); +x_217 = lean_ctor_get(x_215, 1); +lean_inc(x_217); +lean_dec(x_215); +x_218 = lean_ctor_get(x_7, 2); +lean_inc(x_218); lean_dec(x_7); lean_inc(x_19); lean_inc(x_18); @@ -5338,171 +5525,171 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_250); +lean_inc(x_218); lean_inc(x_1); -x_251 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_250, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_248); -if (lean_obj_tag(x_251) == 0) +x_219 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_218, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_217); +if (lean_obj_tag(x_219) == 0) { -lean_object* x_252; lean_object* x_253; lean_object* x_254; uint8_t x_255; -x_252 = lean_ctor_get(x_251, 1); -lean_inc(x_252); -lean_dec(x_251); +lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; +x_220 = lean_ctor_get(x_219, 1); +lean_inc(x_220); +lean_dec(x_219); lean_inc(x_8); -x_253 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_252); -x_254 = lean_ctor_get(x_253, 0); -lean_inc(x_254); -x_255 = lean_unbox(x_254); -lean_dec(x_254); -if (x_255 == 0) -{ -lean_object* x_256; lean_object* x_257; lean_object* x_258; -lean_dec(x_249); +x_221 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_220); +x_222 = lean_ctor_get(x_221, 0); +lean_inc(x_222); +x_223 = lean_unbox(x_222); +lean_dec(x_222); +if (x_223 == 0) +{ +lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_dec(x_8); lean_dec(x_1); -x_256 = lean_ctor_get(x_253, 1); -lean_inc(x_256); -lean_dec(x_253); -x_257 = lean_box(0); -x_258 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_247, x_241, x_9, x_24, x_243, x_242, x_250, x_10, x_245, x_244, x_257, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_256); -lean_dec(x_243); -return x_258; -} -else -{ -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; -x_259 = lean_ctor_get(x_253, 1); -lean_inc(x_259); -if (lean_is_exclusive(x_253)) { - lean_ctor_release(x_253, 0); - lean_ctor_release(x_253, 1); - x_260 = x_253; +x_224 = lean_ctor_get(x_221, 1); +lean_inc(x_224); +lean_dec(x_221); +x_225 = lean_box(0); +x_226 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_216, x_210, x_9, x_24, x_212, x_211, x_218, x_10, x_214, x_213, x_225, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_224); +lean_dec(x_212); +return x_226; +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; +x_227 = lean_ctor_get(x_221, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_221)) { + lean_ctor_release(x_221, 0); + lean_ctor_release(x_221, 1); + x_228 = x_221; } else { - lean_dec_ref(x_253); - x_260 = lean_box(0); + lean_dec_ref(x_221); + x_228 = lean_box(0); } -x_261 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_259); -x_262 = lean_ctor_get(x_261, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_261, 1); -lean_inc(x_263); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - x_264 = x_261; +x_229 = l_Lean_Meta_Grind_updateLastTag(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_227); +if (lean_obj_tag(x_229) == 0) +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_230 = lean_ctor_get(x_229, 1); +lean_inc(x_230); +lean_dec(x_229); +x_231 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_230); +x_232 = lean_ctor_get(x_231, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_231, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_231)) { + lean_ctor_release(x_231, 0); + lean_ctor_release(x_231, 1); + x_234 = x_231; } else { - lean_dec_ref(x_261); - x_264 = lean_box(0); + lean_dec_ref(x_231); + x_234 = lean_box(0); } -x_265 = l_Lean_Meta_Grind_ppENodeRef(x_250, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_263); -x_266 = lean_ctor_get(x_265, 0); -lean_inc(x_266); -x_267 = lean_ctor_get(x_265, 1); -lean_inc(x_267); -if (lean_is_exclusive(x_265)) { - lean_ctor_release(x_265, 0); - lean_ctor_release(x_265, 1); - x_268 = x_265; +x_235 = l_Lean_Meta_Grind_ppENodeRef(x_218, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_233); +x_236 = lean_ctor_get(x_235, 0); +lean_inc(x_236); +x_237 = lean_ctor_get(x_235, 1); +lean_inc(x_237); +if (lean_is_exclusive(x_235)) { + lean_ctor_release(x_235, 0); + lean_ctor_release(x_235, 1); + x_238 = x_235; } else { - lean_dec_ref(x_265); - x_268 = lean_box(0); + lean_dec_ref(x_235); + x_238 = lean_box(0); } -x_269 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_267); -if (lean_obj_tag(x_269) == 0) +x_239 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_237); +if (lean_obj_tag(x_239) == 0) { -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; -x_270 = lean_ctor_get(x_269, 0); -lean_inc(x_270); -x_271 = lean_ctor_get(x_269, 1); -lean_inc(x_271); -lean_dec(x_269); -x_272 = l_Lean_Meta_Grind_ppENodeRef(x_270, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_271); -lean_dec(x_270); -x_273 = lean_ctor_get(x_272, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_272, 1); -lean_inc(x_274); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - lean_ctor_release(x_272, 1); - x_275 = x_272; +lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_240 = lean_ctor_get(x_239, 0); +lean_inc(x_240); +x_241 = lean_ctor_get(x_239, 1); +lean_inc(x_241); +lean_dec(x_239); +x_242 = l_Lean_Meta_Grind_ppENodeRef(x_240, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_241); +lean_dec(x_240); +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); +lean_inc(x_244); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_245 = x_242; } else { - lean_dec_ref(x_272); - x_275 = lean_box(0); + lean_dec_ref(x_242); + x_245 = lean_box(0); } -x_276 = l_Lean_MessageData_ofFormat(x_262); -x_277 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_275)) { - x_278 = lean_alloc_ctor(7, 2, 0); +x_246 = l_Lean_MessageData_ofFormat(x_232); +x_247 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_245)) { + x_248 = lean_alloc_ctor(7, 2, 0); } else { - x_278 = x_275; - lean_ctor_set_tag(x_278, 7); + x_248 = x_245; + lean_ctor_set_tag(x_248, 7); } -lean_ctor_set(x_278, 0, x_277); -lean_ctor_set(x_278, 1, x_276); -x_279 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_268)) { - x_280 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_248, 0, x_247); +lean_ctor_set(x_248, 1, x_246); +x_249 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +if (lean_is_scalar(x_238)) { + x_250 = lean_alloc_ctor(7, 2, 0); } else { - x_280 = x_268; - lean_ctor_set_tag(x_280, 7); + x_250 = x_238; + lean_ctor_set_tag(x_250, 7); } -lean_ctor_set(x_280, 0, x_278); -lean_ctor_set(x_280, 1, x_279); -x_281 = l_Lean_MessageData_ofFormat(x_266); -if (lean_is_scalar(x_264)) { - x_282 = lean_alloc_ctor(7, 2, 0); -} else { - x_282 = x_264; - lean_ctor_set_tag(x_282, 7); -} -lean_ctor_set(x_282, 0, x_280); -lean_ctor_set(x_282, 1, x_281); -x_283 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -if (lean_is_scalar(x_260)) { - x_284 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_250, 0, x_248); +lean_ctor_set(x_250, 1, x_249); +x_251 = l_Lean_MessageData_ofFormat(x_236); +if (lean_is_scalar(x_234)) { + x_252 = lean_alloc_ctor(7, 2, 0); } else { - x_284 = x_260; - lean_ctor_set_tag(x_284, 7); -} -lean_ctor_set(x_284, 0, x_282); -lean_ctor_set(x_284, 1, x_283); -x_285 = l_Lean_MessageData_ofFormat(x_273); -if (lean_is_scalar(x_249)) { - x_286 = lean_alloc_ctor(7, 2, 0); + x_252 = x_234; + lean_ctor_set_tag(x_252, 7); +} +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +x_253 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +if (lean_is_scalar(x_228)) { + x_254 = lean_alloc_ctor(7, 2, 0); } else { - x_286 = x_249; - lean_ctor_set_tag(x_286, 7); -} -lean_ctor_set(x_286, 0, x_284); -lean_ctor_set(x_286, 1, x_285); -x_287 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_287, 0, x_286); -lean_ctor_set(x_287, 1, x_277); -x_288 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_287, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_274); -x_289 = lean_ctor_get(x_288, 0); -lean_inc(x_289); -x_290 = lean_ctor_get(x_288, 1); -lean_inc(x_290); -lean_dec(x_288); -x_291 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_247, x_241, x_9, x_24, x_243, x_242, x_250, x_10, x_245, x_244, x_289, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_290); -lean_dec(x_289); -lean_dec(x_243); -return x_291; -} -else -{ -lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; -lean_dec(x_268); -lean_dec(x_266); -lean_dec(x_264); -lean_dec(x_262); -lean_dec(x_260); -lean_dec(x_250); -lean_dec(x_249); -lean_dec(x_247); -lean_dec(x_243); -lean_dec(x_242); -lean_dec(x_241); + x_254 = x_228; + lean_ctor_set_tag(x_254, 7); +} +lean_ctor_set(x_254, 0, x_252); +lean_ctor_set(x_254, 1, x_253); +x_255 = l_Lean_MessageData_ofFormat(x_243); +x_256 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +x_257 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_257, 0, x_256); +lean_ctor_set(x_257, 1, x_247); +x_258 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_257, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_244); +x_259 = lean_ctor_get(x_258, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_258, 1); +lean_inc(x_260); +lean_dec(x_258); +x_261 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_216, x_210, x_9, x_24, x_212, x_211, x_218, x_10, x_214, x_213, x_259, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_260); +lean_dec(x_259); +lean_dec(x_212); +return x_261; +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; +lean_dec(x_238); +lean_dec(x_236); +lean_dec(x_234); +lean_dec(x_232); +lean_dec(x_228); +lean_dec(x_218); +lean_dec(x_216); +lean_dec(x_212); +lean_dec(x_211); +lean_dec(x_210); lean_dec(x_24); lean_dec(x_19); lean_dec(x_18); @@ -5514,38 +5701,37 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_292 = lean_ctor_get(x_269, 0); -lean_inc(x_292); -x_293 = lean_ctor_get(x_269, 1); -lean_inc(x_293); -if (lean_is_exclusive(x_269)) { - lean_ctor_release(x_269, 0); - lean_ctor_release(x_269, 1); - x_294 = x_269; +x_262 = lean_ctor_get(x_239, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_239, 1); +lean_inc(x_263); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_264 = x_239; } else { - lean_dec_ref(x_269); - x_294 = lean_box(0); + lean_dec_ref(x_239); + x_264 = lean_box(0); } -if (lean_is_scalar(x_294)) { - x_295 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_264)) { + x_265 = lean_alloc_ctor(1, 2, 0); } else { - x_295 = x_294; -} -lean_ctor_set(x_295, 0, x_292); -lean_ctor_set(x_295, 1, x_293); -return x_295; + x_265 = x_264; } +lean_ctor_set(x_265, 0, x_262); +lean_ctor_set(x_265, 1, x_263); +return x_265; } } else { -lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; -lean_dec(x_250); -lean_dec(x_249); -lean_dec(x_247); -lean_dec(x_243); -lean_dec(x_242); -lean_dec(x_241); +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +lean_dec(x_228); +lean_dec(x_218); +lean_dec(x_216); +lean_dec(x_212); +lean_dec(x_211); +lean_dec(x_210); lean_dec(x_24); lean_dec(x_19); lean_dec(x_18); @@ -5558,110 +5744,190 @@ lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_1); -x_296 = lean_ctor_get(x_251, 0); -lean_inc(x_296); -x_297 = lean_ctor_get(x_251, 1); -lean_inc(x_297); -if (lean_is_exclusive(x_251)) { - lean_ctor_release(x_251, 0); - lean_ctor_release(x_251, 1); - x_298 = x_251; +x_266 = lean_ctor_get(x_229, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_229, 1); +lean_inc(x_267); +if (lean_is_exclusive(x_229)) { + lean_ctor_release(x_229, 0); + lean_ctor_release(x_229, 1); + x_268 = x_229; } else { - lean_dec_ref(x_251); - x_298 = lean_box(0); + lean_dec_ref(x_229); + x_268 = lean_box(0); } -if (lean_is_scalar(x_298)) { - x_299 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_268)) { + x_269 = lean_alloc_ctor(1, 2, 0); } else { - x_299 = x_298; + x_269 = x_268; } -lean_ctor_set(x_299, 0, x_296); -lean_ctor_set(x_299, 1, x_297); -return x_299; +lean_ctor_set(x_269, 0, x_266); +lean_ctor_set(x_269, 1, x_267); +return x_269; } } } else { -lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; uint8_t x_305; uint8_t x_306; uint8_t x_307; uint8_t x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; uint8_t x_321; uint8_t x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; -x_300 = lean_ctor_get(x_2, 0); -x_301 = lean_ctor_get(x_2, 1); -x_302 = lean_ctor_get(x_2, 2); -x_303 = lean_ctor_get(x_2, 3); -x_304 = lean_ctor_get(x_2, 6); -x_305 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 1); -x_306 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 2); -x_307 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 3); -x_308 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 4); -x_309 = lean_ctor_get(x_2, 7); -x_310 = lean_ctor_get(x_2, 8); -x_311 = lean_ctor_get(x_2, 9); -lean_inc(x_311); -lean_inc(x_310); -lean_inc(x_309); -lean_inc(x_304); -lean_inc(x_303); -lean_inc(x_302); -lean_inc(x_301); -lean_inc(x_300); +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +lean_dec(x_218); +lean_dec(x_216); +lean_dec(x_212); +lean_dec(x_211); +lean_dec(x_210); +lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_270 = lean_ctor_get(x_219, 0); +lean_inc(x_270); +x_271 = lean_ctor_get(x_219, 1); +lean_inc(x_271); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_272 = x_219; +} else { + lean_dec_ref(x_219); + x_272 = lean_box(0); +} +if (lean_is_scalar(x_272)) { + x_273 = lean_alloc_ctor(1, 2, 0); +} else { + x_273 = x_272; +} +lean_ctor_set(x_273, 0, x_270); +lean_ctor_set(x_273, 1, x_271); +return x_273; +} +} +else +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; +lean_dec(x_212); +lean_dec(x_211); +lean_dec(x_210); +lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_1); +x_274 = lean_ctor_get(x_215, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_215, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_215)) { + lean_ctor_release(x_215, 0); + lean_ctor_release(x_215, 1); + x_276 = x_215; +} else { + lean_dec_ref(x_215); + x_276 = lean_box(0); +} +if (lean_is_scalar(x_276)) { + x_277 = lean_alloc_ctor(1, 2, 0); +} else { + x_277 = x_276; +} +lean_ctor_set(x_277, 0, x_274); +lean_ctor_set(x_277, 1, x_275); +return x_277; +} +} +} +else +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; uint8_t x_284; uint8_t x_285; uint8_t x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; uint8_t x_299; uint8_t x_300; lean_object* x_301; +x_278 = lean_ctor_get(x_2, 0); +x_279 = lean_ctor_get(x_2, 1); +x_280 = lean_ctor_get(x_2, 2); +x_281 = lean_ctor_get(x_2, 3); +x_282 = lean_ctor_get(x_2, 6); +x_283 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 1); +x_284 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 2); +x_285 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 3); +x_286 = lean_ctor_get_uint8(x_2, sizeof(void*)*10 + 4); +x_287 = lean_ctor_get(x_2, 7); +x_288 = lean_ctor_get(x_2, 8); +x_289 = lean_ctor_get(x_2, 9); +lean_inc(x_289); +lean_inc(x_288); +lean_inc(x_287); +lean_inc(x_282); +lean_inc(x_281); +lean_inc(x_280); +lean_inc(x_279); +lean_inc(x_278); lean_dec(x_2); -x_312 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_312, 0, x_3); -x_313 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_313, 0, x_4); -lean_inc(x_302); -x_314 = lean_alloc_ctor(0, 10, 5); -lean_ctor_set(x_314, 0, x_300); -lean_ctor_set(x_314, 1, x_301); -lean_ctor_set(x_314, 2, x_302); -lean_ctor_set(x_314, 3, x_303); -lean_ctor_set(x_314, 4, x_312); -lean_ctor_set(x_314, 5, x_313); -lean_ctor_set(x_314, 6, x_304); -lean_ctor_set(x_314, 7, x_309); -lean_ctor_set(x_314, 8, x_310); -lean_ctor_set(x_314, 9, x_311); -lean_ctor_set_uint8(x_314, sizeof(void*)*10, x_5); -lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 1, x_305); -lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 2, x_306); -lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 3, x_307); -lean_ctor_set_uint8(x_314, sizeof(void*)*10 + 4, x_308); +x_290 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_290, 0, x_3); +x_291 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_291, 0, x_4); +lean_inc(x_280); +x_292 = lean_alloc_ctor(0, 10, 5); +lean_ctor_set(x_292, 0, x_278); +lean_ctor_set(x_292, 1, x_279); +lean_ctor_set(x_292, 2, x_280); +lean_ctor_set(x_292, 3, x_281); +lean_ctor_set(x_292, 4, x_290); +lean_ctor_set(x_292, 5, x_291); +lean_ctor_set(x_292, 6, x_282); +lean_ctor_set(x_292, 7, x_287); +lean_ctor_set(x_292, 8, x_288); +lean_ctor_set(x_292, 9, x_289); +lean_ctor_set_uint8(x_292, sizeof(void*)*10, x_5); +lean_ctor_set_uint8(x_292, sizeof(void*)*10 + 1, x_283); +lean_ctor_set_uint8(x_292, sizeof(void*)*10 + 2, x_284); +lean_ctor_set_uint8(x_292, sizeof(void*)*10 + 3, x_285); +lean_ctor_set_uint8(x_292, sizeof(void*)*10 + 4, x_286); lean_inc(x_1); -x_315 = l_Lean_Meta_Grind_setENode(x_1, x_314, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_22); -x_316 = lean_ctor_get(x_315, 1); -lean_inc(x_316); -if (lean_is_exclusive(x_315)) { - lean_ctor_release(x_315, 0); - lean_ctor_release(x_315, 1); - x_317 = x_315; +x_293 = l_Lean_Meta_Grind_setENode(x_1, x_292, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_22); +x_294 = lean_ctor_get(x_293, 1); +lean_inc(x_294); +if (lean_is_exclusive(x_293)) { + lean_ctor_release(x_293, 0); + lean_ctor_release(x_293, 1); + x_295 = x_293; } else { - lean_dec_ref(x_315); - x_317 = lean_box(0); + lean_dec_ref(x_293); + x_295 = lean_box(0); } -x_318 = lean_ctor_get(x_6, 0); -lean_inc(x_318); -x_319 = lean_ctor_get(x_6, 1); -lean_inc(x_319); -x_320 = lean_ctor_get(x_6, 6); -lean_inc(x_320); -x_321 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); -x_322 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); +x_296 = lean_ctor_get(x_6, 0); +lean_inc(x_296); +x_297 = lean_ctor_get(x_6, 1); +lean_inc(x_297); +x_298 = lean_ctor_get(x_6, 6); +lean_inc(x_298); +x_299 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 3); +x_300 = lean_ctor_get_uint8(x_6, sizeof(void*)*10 + 4); lean_dec(x_6); -x_323 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_318, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_316); -x_324 = lean_ctor_get(x_323, 0); -lean_inc(x_324); -x_325 = lean_ctor_get(x_323, 1); -lean_inc(x_325); -if (lean_is_exclusive(x_323)) { - lean_ctor_release(x_323, 0); - lean_ctor_release(x_323, 1); - x_326 = x_323; -} else { - lean_dec_ref(x_323); - x_326 = lean_box(0); -} -x_327 = lean_ctor_get(x_7, 2); -lean_inc(x_327); +x_301 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents(x_296, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_294); +if (lean_obj_tag(x_301) == 0) +{ +lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; +x_302 = lean_ctor_get(x_301, 0); +lean_inc(x_302); +x_303 = lean_ctor_get(x_301, 1); +lean_inc(x_303); +lean_dec(x_301); +x_304 = lean_ctor_get(x_7, 2); +lean_inc(x_304); lean_dec(x_7); lean_inc(x_19); lean_inc(x_18); @@ -5671,179 +5937,222 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_327); +lean_inc(x_304); lean_inc(x_1); -x_328 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_327, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_325); -if (lean_obj_tag(x_328) == 0) +x_305 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_updateRoots_loop(x_1, x_304, x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_303); +if (lean_obj_tag(x_305) == 0) { -lean_object* x_329; lean_object* x_330; lean_object* x_331; uint8_t x_332; -x_329 = lean_ctor_get(x_328, 1); -lean_inc(x_329); -lean_dec(x_328); +lean_object* x_306; lean_object* x_307; lean_object* x_308; uint8_t x_309; +x_306 = lean_ctor_get(x_305, 1); +lean_inc(x_306); +lean_dec(x_305); lean_inc(x_8); -x_330 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_329); -x_331 = lean_ctor_get(x_330, 0); -lean_inc(x_331); -x_332 = lean_unbox(x_331); -lean_dec(x_331); -if (x_332 == 0) -{ -lean_object* x_333; lean_object* x_334; lean_object* x_335; -lean_dec(x_326); -lean_dec(x_317); +x_307 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_8, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_306); +x_308 = lean_ctor_get(x_307, 0); +lean_inc(x_308); +x_309 = lean_unbox(x_308); +lean_dec(x_308); +if (x_309 == 0) +{ +lean_object* x_310; lean_object* x_311; lean_object* x_312; +lean_dec(x_295); lean_dec(x_8); lean_dec(x_1); -x_333 = lean_ctor_get(x_330, 1); -lean_inc(x_333); -lean_dec(x_330); -x_334 = lean_box(0); -x_335 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_324, x_318, x_9, x_302, x_320, x_319, x_327, x_10, x_322, x_321, x_334, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_333); -lean_dec(x_320); -return x_335; +x_310 = lean_ctor_get(x_307, 1); +lean_inc(x_310); +lean_dec(x_307); +x_311 = lean_box(0); +x_312 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_302, x_296, x_9, x_280, x_298, x_297, x_304, x_10, x_300, x_299, x_311, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_310); +lean_dec(x_298); +return x_312; } else { -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; -x_336 = lean_ctor_get(x_330, 1); -lean_inc(x_336); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_337 = x_330; +lean_object* x_313; lean_object* x_314; lean_object* x_315; +x_313 = lean_ctor_get(x_307, 1); +lean_inc(x_313); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_314 = x_307; } else { - lean_dec_ref(x_330); - x_337 = lean_box(0); -} -x_338 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_336); -x_339 = lean_ctor_get(x_338, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_338, 1); -lean_inc(x_340); -if (lean_is_exclusive(x_338)) { - lean_ctor_release(x_338, 0); - lean_ctor_release(x_338, 1); - x_341 = x_338; -} else { - lean_dec_ref(x_338); - x_341 = lean_box(0); -} -x_342 = l_Lean_Meta_Grind_ppENodeRef(x_327, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_340); -x_343 = lean_ctor_get(x_342, 0); -lean_inc(x_343); -x_344 = lean_ctor_get(x_342, 1); -lean_inc(x_344); -if (lean_is_exclusive(x_342)) { - lean_ctor_release(x_342, 0); - lean_ctor_release(x_342, 1); - x_345 = x_342; -} else { - lean_dec_ref(x_342); - x_345 = lean_box(0); + lean_dec_ref(x_307); + x_314 = lean_box(0); } -x_346 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_344); -if (lean_obj_tag(x_346) == 0) +x_315 = l_Lean_Meta_Grind_updateLastTag(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_313); +if (lean_obj_tag(x_315) == 0) { -lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; -x_347 = lean_ctor_get(x_346, 0); -lean_inc(x_347); -x_348 = lean_ctor_get(x_346, 1); -lean_inc(x_348); -lean_dec(x_346); -x_349 = l_Lean_Meta_Grind_ppENodeRef(x_347, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_348); -lean_dec(x_347); -x_350 = lean_ctor_get(x_349, 0); -lean_inc(x_350); -x_351 = lean_ctor_get(x_349, 1); -lean_inc(x_351); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - x_352 = x_349; +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; +x_316 = lean_ctor_get(x_315, 1); +lean_inc(x_316); +lean_dec(x_315); +x_317 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_316); +x_318 = lean_ctor_get(x_317, 0); +lean_inc(x_318); +x_319 = lean_ctor_get(x_317, 1); +lean_inc(x_319); +if (lean_is_exclusive(x_317)) { + lean_ctor_release(x_317, 0); + lean_ctor_release(x_317, 1); + x_320 = x_317; +} else { + lean_dec_ref(x_317); + x_320 = lean_box(0); +} +x_321 = l_Lean_Meta_Grind_ppENodeRef(x_304, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_319); +x_322 = lean_ctor_get(x_321, 0); +lean_inc(x_322); +x_323 = lean_ctor_get(x_321, 1); +lean_inc(x_323); +if (lean_is_exclusive(x_321)) { + lean_ctor_release(x_321, 0); + lean_ctor_release(x_321, 1); + x_324 = x_321; } else { - lean_dec_ref(x_349); - x_352 = lean_box(0); + lean_dec_ref(x_321); + x_324 = lean_box(0); } -x_353 = l_Lean_MessageData_ofFormat(x_339); -x_354 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_352)) { - x_355 = lean_alloc_ctor(7, 2, 0); +x_325 = l_Lean_Meta_Grind_getRoot(x_1, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_323); +if (lean_obj_tag(x_325) == 0) +{ +lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; +x_326 = lean_ctor_get(x_325, 0); +lean_inc(x_326); +x_327 = lean_ctor_get(x_325, 1); +lean_inc(x_327); +lean_dec(x_325); +x_328 = l_Lean_Meta_Grind_ppENodeRef(x_326, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_327); +lean_dec(x_326); +x_329 = lean_ctor_get(x_328, 0); +lean_inc(x_329); +x_330 = lean_ctor_get(x_328, 1); +lean_inc(x_330); +if (lean_is_exclusive(x_328)) { + lean_ctor_release(x_328, 0); + lean_ctor_release(x_328, 1); + x_331 = x_328; } else { - x_355 = x_352; - lean_ctor_set_tag(x_355, 7); + lean_dec_ref(x_328); + x_331 = lean_box(0); } -lean_ctor_set(x_355, 0, x_354); -lean_ctor_set(x_355, 1, x_353); -x_356 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; -if (lean_is_scalar(x_345)) { - x_357 = lean_alloc_ctor(7, 2, 0); +x_332 = l_Lean_MessageData_ofFormat(x_318); +x_333 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_331)) { + x_334 = lean_alloc_ctor(7, 2, 0); } else { - x_357 = x_345; - lean_ctor_set_tag(x_357, 7); -} -lean_ctor_set(x_357, 0, x_355); -lean_ctor_set(x_357, 1, x_356); -x_358 = l_Lean_MessageData_ofFormat(x_343); -if (lean_is_scalar(x_341)) { - x_359 = lean_alloc_ctor(7, 2, 0); + x_334 = x_331; + lean_ctor_set_tag(x_334, 7); +} +lean_ctor_set(x_334, 0, x_333); +lean_ctor_set(x_334, 1, x_332); +x_335 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__2; +if (lean_is_scalar(x_324)) { + x_336 = lean_alloc_ctor(7, 2, 0); } else { - x_359 = x_341; - lean_ctor_set_tag(x_359, 7); -} -lean_ctor_set(x_359, 0, x_357); -lean_ctor_set(x_359, 1, x_358); -x_360 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; -if (lean_is_scalar(x_337)) { - x_361 = lean_alloc_ctor(7, 2, 0); + x_336 = x_324; + lean_ctor_set_tag(x_336, 7); +} +lean_ctor_set(x_336, 0, x_334); +lean_ctor_set(x_336, 1, x_335); +x_337 = l_Lean_MessageData_ofFormat(x_322); +if (lean_is_scalar(x_320)) { + x_338 = lean_alloc_ctor(7, 2, 0); } else { - x_361 = x_337; - lean_ctor_set_tag(x_361, 7); -} -lean_ctor_set(x_361, 0, x_359); -lean_ctor_set(x_361, 1, x_360); -x_362 = l_Lean_MessageData_ofFormat(x_350); -if (lean_is_scalar(x_326)) { - x_363 = lean_alloc_ctor(7, 2, 0); + x_338 = x_320; + lean_ctor_set_tag(x_338, 7); +} +lean_ctor_set(x_338, 0, x_336); +lean_ctor_set(x_338, 1, x_337); +x_339 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3___closed__4; +if (lean_is_scalar(x_314)) { + x_340 = lean_alloc_ctor(7, 2, 0); } else { - x_363 = x_326; - lean_ctor_set_tag(x_363, 7); -} -lean_ctor_set(x_363, 0, x_361); -lean_ctor_set(x_363, 1, x_362); -if (lean_is_scalar(x_317)) { - x_364 = lean_alloc_ctor(7, 2, 0); + x_340 = x_314; + lean_ctor_set_tag(x_340, 7); +} +lean_ctor_set(x_340, 0, x_338); +lean_ctor_set(x_340, 1, x_339); +x_341 = l_Lean_MessageData_ofFormat(x_329); +if (lean_is_scalar(x_295)) { + x_342 = lean_alloc_ctor(7, 2, 0); } else { - x_364 = x_317; - lean_ctor_set_tag(x_364, 7); -} -lean_ctor_set(x_364, 0, x_363); -lean_ctor_set(x_364, 1, x_354); -x_365 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_8, x_364, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_351); -x_366 = lean_ctor_get(x_365, 0); -lean_inc(x_366); -x_367 = lean_ctor_get(x_365, 1); -lean_inc(x_367); -lean_dec(x_365); -x_368 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_324, x_318, x_9, x_302, x_320, x_319, x_327, x_10, x_322, x_321, x_366, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_367); -lean_dec(x_366); -lean_dec(x_320); -return x_368; + x_342 = x_295; + lean_ctor_set_tag(x_342, 7); +} +lean_ctor_set(x_342, 0, x_340); +lean_ctor_set(x_342, 1, x_341); +x_343 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_343, 0, x_342); +lean_ctor_set(x_343, 1, x_333); +x_344 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_8, x_343, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_330); +x_345 = lean_ctor_get(x_344, 0); +lean_inc(x_345); +x_346 = lean_ctor_get(x_344, 1); +lean_inc(x_346); +lean_dec(x_344); +x_347 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__2(x_302, x_296, x_9, x_280, x_298, x_297, x_304, x_10, x_300, x_299, x_345, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_346); +lean_dec(x_345); +lean_dec(x_298); +return x_347; } else { -lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; -lean_dec(x_345); -lean_dec(x_343); -lean_dec(x_341); -lean_dec(x_339); -lean_dec(x_337); -lean_dec(x_327); -lean_dec(x_326); +lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_dec(x_324); +lean_dec(x_322); lean_dec(x_320); -lean_dec(x_319); lean_dec(x_318); -lean_dec(x_317); +lean_dec(x_314); +lean_dec(x_304); +lean_dec(x_302); +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_280); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +x_348 = lean_ctor_get(x_325, 0); +lean_inc(x_348); +x_349 = lean_ctor_get(x_325, 1); +lean_inc(x_349); +if (lean_is_exclusive(x_325)) { + lean_ctor_release(x_325, 0); + lean_ctor_release(x_325, 1); + x_350 = x_325; +} else { + lean_dec_ref(x_325); + x_350 = lean_box(0); +} +if (lean_is_scalar(x_350)) { + x_351 = lean_alloc_ctor(1, 2, 0); +} else { + x_351 = x_350; +} +lean_ctor_set(x_351, 0, x_348); +lean_ctor_set(x_351, 1, x_349); +return x_351; +} +} +else +{ +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_314); +lean_dec(x_304); lean_dec(x_302); +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_280); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -5854,40 +6163,40 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); -x_369 = lean_ctor_get(x_346, 0); -lean_inc(x_369); -x_370 = lean_ctor_get(x_346, 1); -lean_inc(x_370); -if (lean_is_exclusive(x_346)) { - lean_ctor_release(x_346, 0); - lean_ctor_release(x_346, 1); - x_371 = x_346; +lean_dec(x_1); +x_352 = lean_ctor_get(x_315, 0); +lean_inc(x_352); +x_353 = lean_ctor_get(x_315, 1); +lean_inc(x_353); +if (lean_is_exclusive(x_315)) { + lean_ctor_release(x_315, 0); + lean_ctor_release(x_315, 1); + x_354 = x_315; } else { - lean_dec_ref(x_346); - x_371 = lean_box(0); + lean_dec_ref(x_315); + x_354 = lean_box(0); } -if (lean_is_scalar(x_371)) { - x_372 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_354)) { + x_355 = lean_alloc_ctor(1, 2, 0); } else { - x_372 = x_371; + x_355 = x_354; } -lean_ctor_set(x_372, 0, x_369); -lean_ctor_set(x_372, 1, x_370); -return x_372; +lean_ctor_set(x_355, 0, x_352); +lean_ctor_set(x_355, 1, x_353); +return x_355; } } } else { -lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; -lean_dec(x_327); -lean_dec(x_326); -lean_dec(x_324); -lean_dec(x_320); -lean_dec(x_319); -lean_dec(x_318); -lean_dec(x_317); +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; +lean_dec(x_304); lean_dec(x_302); +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_280); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -5899,32 +6208,74 @@ lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_1); -x_373 = lean_ctor_get(x_328, 0); -lean_inc(x_373); -x_374 = lean_ctor_get(x_328, 1); -lean_inc(x_374); -if (lean_is_exclusive(x_328)) { - lean_ctor_release(x_328, 0); - lean_ctor_release(x_328, 1); - x_375 = x_328; +x_356 = lean_ctor_get(x_305, 0); +lean_inc(x_356); +x_357 = lean_ctor_get(x_305, 1); +lean_inc(x_357); +if (lean_is_exclusive(x_305)) { + lean_ctor_release(x_305, 0); + lean_ctor_release(x_305, 1); + x_358 = x_305; } else { - lean_dec_ref(x_328); - x_375 = lean_box(0); + lean_dec_ref(x_305); + x_358 = lean_box(0); +} +if (lean_is_scalar(x_358)) { + x_359 = lean_alloc_ctor(1, 2, 0); +} else { + x_359 = x_358; +} +lean_ctor_set(x_359, 0, x_356); +lean_ctor_set(x_359, 1, x_357); +return x_359; +} +} +else +{ +lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; +lean_dec(x_298); +lean_dec(x_297); +lean_dec(x_296); +lean_dec(x_295); +lean_dec(x_280); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_1); +x_360 = lean_ctor_get(x_301, 0); +lean_inc(x_360); +x_361 = lean_ctor_get(x_301, 1); +lean_inc(x_361); +if (lean_is_exclusive(x_301)) { + lean_ctor_release(x_301, 0); + lean_ctor_release(x_301, 1); + x_362 = x_301; +} else { + lean_dec_ref(x_301); + x_362 = lean_box(0); } -if (lean_is_scalar(x_375)) { - x_376 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_362)) { + x_363 = lean_alloc_ctor(1, 2, 0); } else { - x_376 = x_375; + x_363 = x_362; } -lean_ctor_set(x_376, 0, x_373); -lean_ctor_set(x_376, 1, x_374); -return x_376; +lean_ctor_set(x_363, 0, x_360); +lean_ctor_set(x_363, 1, x_361); +return x_363; } } } else { -uint8_t x_377; +uint8_t x_364; lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -5941,23 +6292,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_377 = !lean_is_exclusive(x_21); -if (x_377 == 0) +x_364 = !lean_is_exclusive(x_21); +if (x_364 == 0) { return x_21; } else { -lean_object* x_378; lean_object* x_379; lean_object* x_380; -x_378 = lean_ctor_get(x_21, 0); -x_379 = lean_ctor_get(x_21, 1); -lean_inc(x_379); -lean_inc(x_378); +lean_object* x_365; lean_object* x_366; lean_object* x_367; +x_365 = lean_ctor_get(x_21, 0); +x_366 = lean_ctor_get(x_21, 1); +lean_inc(x_366); +lean_inc(x_365); lean_dec(x_21); -x_380 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_380, 0, x_378); -lean_ctor_set(x_380, 1, x_379); -return x_380; +x_367 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_367, 0, x_365); +lean_ctor_set(x_367, 1, x_366); +return x_367; } } } @@ -6011,7 +6362,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr { lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; x_19 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_20 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_19, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +x_20 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_19, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); x_22 = lean_unbox(x_21); @@ -6032,209 +6383,302 @@ uint8_t x_26; x_26 = !lean_is_exclusive(x_20); if (x_26 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_20, 1); x_28 = lean_ctor_get(x_20, 0); lean_dec(x_28); -x_29 = l_Lean_Meta_Grind_ppENodeRef(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_27); -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +x_29 = l_Lean_Meta_Grind_updateLastTag(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_31 = lean_ctor_get(x_29, 0); -x_32 = lean_ctor_get(x_29, 1); -x_33 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_32); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_Lean_Meta_Grind_ppENodeRef(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_30); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_35 = lean_ctor_get(x_33, 0); -x_36 = lean_ctor_get(x_33, 1); -x_37 = l_Lean_MessageData_ofFormat(x_31); -x_38 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -lean_ctor_set_tag(x_33, 7); -lean_ctor_set(x_33, 1, x_37); -lean_ctor_set(x_33, 0, x_38); -x_39 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; -lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_39); -lean_ctor_set(x_29, 0, x_33); -x_40 = l_Lean_MessageData_ofFormat(x_35); +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +x_35 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_34); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_37 = lean_ctor_get(x_35, 0); +x_38 = lean_ctor_get(x_35, 1); +x_39 = l_Lean_MessageData_ofFormat(x_33); +x_40 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_39); +lean_ctor_set(x_35, 0, x_40); +x_41 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +lean_ctor_set_tag(x_31, 7); +lean_ctor_set(x_31, 1, x_41); +lean_ctor_set(x_31, 0, x_35); +x_42 = l_Lean_MessageData_ofFormat(x_37); lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_40); -lean_ctor_set(x_20, 0, x_29); -x_41 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_20); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_19, x_42, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_36); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_44, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_45); -lean_dec(x_44); -return x_46; +lean_ctor_set(x_20, 1, x_42); +lean_ctor_set(x_20, 0, x_31); +x_43 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_20); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_19, x_44, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_38); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_46, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_47); +lean_dec(x_46); +return x_48; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_47 = lean_ctor_get(x_33, 0); -x_48 = lean_ctor_get(x_33, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_33); -x_49 = l_Lean_MessageData_ofFormat(x_31); -x_50 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_49); -x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; -lean_ctor_set_tag(x_29, 7); -lean_ctor_set(x_29, 1, x_52); -lean_ctor_set(x_29, 0, x_51); -x_53 = l_Lean_MessageData_ofFormat(x_47); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_49 = lean_ctor_get(x_35, 0); +x_50 = lean_ctor_get(x_35, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_35); +x_51 = l_Lean_MessageData_ofFormat(x_33); +x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +x_54 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +lean_ctor_set_tag(x_31, 7); +lean_ctor_set(x_31, 1, x_54); +lean_ctor_set(x_31, 0, x_53); +x_55 = l_Lean_MessageData_ofFormat(x_49); lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_53); -lean_ctor_set(x_20, 0, x_29); -x_54 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_20); -lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_19, x_55, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_48); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_57, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_58); -lean_dec(x_57); -return x_59; +lean_ctor_set(x_20, 1, x_55); +lean_ctor_set(x_20, 0, x_31); +x_56 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_20); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_19, x_57, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_50); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_59, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_60); +lean_dec(x_59); +return x_61; } } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_60 = lean_ctor_get(x_29, 0); -x_61 = lean_ctor_get(x_29, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_29); -x_62 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_61); -x_63 = lean_ctor_get(x_62, 0); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_62 = lean_ctor_get(x_31, 0); +x_63 = lean_ctor_get(x_31, 1); lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_65 = x_62; +lean_inc(x_62); +lean_dec(x_31); +x_64 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_63); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_67 = x_64; } else { - lean_dec_ref(x_62); - x_65 = lean_box(0); + lean_dec_ref(x_64); + x_67 = lean_box(0); } -x_66 = l_Lean_MessageData_ofFormat(x_60); -x_67 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -if (lean_is_scalar(x_65)) { - x_68 = lean_alloc_ctor(7, 2, 0); +x_68 = l_Lean_MessageData_ofFormat(x_62); +x_69 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +if (lean_is_scalar(x_67)) { + x_70 = lean_alloc_ctor(7, 2, 0); } else { - x_68 = x_65; - lean_ctor_set_tag(x_68, 7); -} -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; -x_70 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean_MessageData_ofFormat(x_63); + x_70 = x_67; + lean_ctor_set_tag(x_70, 7); +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +x_71 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_MessageData_ofFormat(x_65); lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_71); -lean_ctor_set(x_20, 0, x_70); -x_72 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_20); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_19, x_73, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_64); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); -lean_inc(x_76); -lean_dec(x_74); -x_77 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_75, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_76); -lean_dec(x_75); -return x_77; +lean_ctor_set(x_20, 1, x_73); +lean_ctor_set(x_20, 0, x_72); +x_74 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_20); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_19, x_75, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_66); +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_77, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_78); +lean_dec(x_77); +return x_79; } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_78 = lean_ctor_get(x_20, 1); -lean_inc(x_78); -lean_dec(x_20); -x_79 = l_Lean_Meta_Grind_ppENodeRef(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_78); -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_82 = x_79; -} else { - lean_dec_ref(x_79); - x_82 = lean_box(0); +uint8_t x_80; +lean_free_object(x_20); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_80 = !lean_is_exclusive(x_29); +if (x_80 == 0) +{ +return x_29; } -x_83 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_81); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); -lean_inc(x_85); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_86 = x_83; -} else { - lean_dec_ref(x_83); - x_86 = lean_box(0); +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_29, 0); +x_82 = lean_ctor_get(x_29, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_29); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; } -x_87 = l_Lean_MessageData_ofFormat(x_80); -x_88 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; -if (lean_is_scalar(x_86)) { - x_89 = lean_alloc_ctor(7, 2, 0); +} +} +else +{ +lean_object* x_84; lean_object* x_85; +x_84 = lean_ctor_get(x_20, 1); +lean_inc(x_84); +lean_dec(x_20); +x_85 = l_Lean_Meta_Grind_updateLastTag(x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_84); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +lean_dec(x_85); +x_87 = l_Lean_Meta_Grind_ppENodeRef(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_86); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; } else { - x_89 = x_86; - lean_ctor_set_tag(x_89, 7); -} -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_87); -x_90 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; -if (lean_is_scalar(x_82)) { - x_91 = lean_alloc_ctor(7, 2, 0); + lean_dec_ref(x_87); + x_90 = lean_box(0); +} +x_91 = l_Lean_Meta_Grind_ppENodeRef(x_4, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_89); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_94 = x_91; } else { - x_91 = x_82; - lean_ctor_set_tag(x_91, 7); + lean_dec_ref(x_91); + x_94 = lean_box(0); +} +x_95 = l_Lean_MessageData_ofFormat(x_88); +x_96 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__3; +if (lean_is_scalar(x_94)) { + x_97 = lean_alloc_ctor(7, 2, 0); +} else { + x_97 = x_94; + lean_ctor_set_tag(x_97, 7); +} +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__5; +if (lean_is_scalar(x_90)) { + x_99 = lean_alloc_ctor(7, 2, 0); +} else { + x_99 = x_90; + lean_ctor_set_tag(x_99, 7); +} +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_MessageData_ofFormat(x_92); +x_101 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_103 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +x_104 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_19, x_103, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_93); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_107 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_105, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_106); +lean_dec(x_105); +return x_107; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_108 = lean_ctor_get(x_85, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_85, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_110 = x_85; +} else { + lean_dec_ref(x_85); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; } -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean_MessageData_ofFormat(x_84); -x_93 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_95 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_19, x_95, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_85); -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___lambda__3(x_3, x_5, x_4, x_1, x_9, x_7, x_6, x_19, x_8, x_2, x_97, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_98); -lean_dec(x_97); -return x_99; } } } @@ -6395,7 +6839,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr { lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_13 = !lean_is_exclusive(x_12); if (x_13 == 0) { @@ -6416,6 +6860,13 @@ return x_19; else { lean_object* x_20; +x_20 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -6424,36 +6875,36 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_20 = l_Lean_Meta_Grind_ppState(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -if (lean_obj_tag(x_20) == 0) +x_22 = l_Lean_Meta_Grind_ppState(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l_Lean_MessageData_ofFormat(x_21); -x_24 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_MessageData_ofFormat(x_23); +x_26 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3; lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_23); -lean_ctor_set(x_12, 0, x_24); -x_25 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_12); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_11, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_apply_10(x_16, x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -return x_30; +lean_ctor_set(x_12, 1, x_25); +lean_ctor_set(x_12, 0, x_26); +x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_12); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_11, x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_24); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_apply_10(x_16, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +return x_32; } else { -uint8_t x_31; +uint8_t x_33; lean_free_object(x_12); lean_dec(x_9); lean_dec(x_8); @@ -6463,48 +6914,87 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_31 = !lean_is_exclusive(x_20); -if (x_31 == 0) +x_33 = !lean_is_exclusive(x_22); +if (x_33 == 0) +{ +return x_22; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_22, 0); +x_35 = lean_ctor_get(x_22, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_22); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_free_object(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_20); +if (x_37 == 0) { return x_20; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_20, 0); -x_33 = lean_ctor_get(x_20, 1); -lean_inc(x_33); -lean_inc(x_32); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_20, 0); +x_39 = lean_ctor_get(x_20, 1); +lean_inc(x_39); +lean_inc(x_38); lean_dec(x_20); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; } } } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_35 = lean_ctor_get(x_12, 0); -x_36 = lean_ctor_get(x_12, 1); -lean_inc(x_36); -lean_inc(x_35); +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = lean_ctor_get(x_12, 0); +x_42 = lean_ctor_get(x_12, 1); +lean_inc(x_42); +lean_inc(x_41); lean_dec(x_12); -x_37 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__1; -x_38 = lean_unbox(x_35); -lean_dec(x_35); -if (x_38 == 0) +x_43 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__1; +x_44 = lean_unbox(x_41); +lean_dec(x_41); +if (x_44 == 0) { -lean_object* x_39; lean_object* x_40; -x_39 = lean_box(0); -x_40 = lean_apply_10(x_37, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_36); -return x_40; +lean_object* x_45; lean_object* x_46; +x_45 = lean_box(0); +x_46 = lean_apply_10(x_43, x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +return x_46; } else { -lean_object* x_41; +lean_object* x_47; +x_47 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -6513,36 +7003,36 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_41 = l_Lean_Meta_Grind_ppState(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_36); -if (lean_obj_tag(x_41) == 0) +x_49 = l_Lean_Meta_Grind_ppState(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_48); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_MessageData_ofFormat(x_42); -x_45 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3; -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_44); -x_47 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_11, x_48, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); lean_dec(x_49); -x_52 = lean_apply_10(x_37, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_51); -return x_52; +x_52 = l_Lean_MessageData_ofFormat(x_50); +x_53 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__2___closed__3; +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); +x_55 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_11, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_51); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_60 = lean_apply_10(x_43, x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_59); +return x_60; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -6551,26 +7041,59 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_53 = lean_ctor_get(x_41, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_41, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - x_55 = x_41; +x_61 = lean_ctor_get(x_49, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_49, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_63 = x_49; } else { - lean_dec_ref(x_41); - x_55 = lean_box(0); + lean_dec_ref(x_49); + x_63 = lean_box(0); } -if (lean_is_scalar(x_55)) { - x_56 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_63)) { + x_64 = lean_alloc_ctor(1, 2, 0); } else { - x_56 = x_55; + x_64 = x_63; } -lean_ctor_set(x_56, 0, x_53); -lean_ctor_set(x_56, 1, x_54); -return x_56; +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_62); +return x_64; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_65 = lean_ctor_get(x_47, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_47, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_67 = x_47; +} else { + lean_dec_ref(x_47); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(1, 2, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; } } } @@ -7446,8 +7969,19 @@ return x_41; } else { -lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; +lean_object* x_42; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); x_42 = l_Lean_Meta_Grind_markAsInconsistent(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; uint8_t x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; x_43 = lean_ctor_get(x_42, 1); lean_inc(x_43); lean_dec(x_42); @@ -7457,27 +7991,114 @@ x_46 = lean_box(0); x_47 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__6(x_19, x_29, x_3, x_4, x_5, x_6, x_1, x_2, x_44, x_45, x_46, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_43); return x_47; } +else +{ +uint8_t x_48; +lean_dec(x_29); +lean_dec(x_19); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_42); +if (x_48 == 0) +{ +return x_42; } else { -lean_object* x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; -lean_dec(x_21); -x_48 = l_Lean_Meta_Grind_markAsInconsistent(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); -x_49 = lean_ctor_get(x_48, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_42, 0); +x_50 = lean_ctor_get(x_42, 1); +lean_inc(x_50); lean_inc(x_49); -lean_dec(x_48); -x_50 = 1; -x_51 = 0; -x_52 = lean_box(0); -x_53 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__6(x_19, x_29, x_3, x_4, x_5, x_6, x_1, x_2, x_50, x_51, x_52, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_49); -return x_53; +lean_dec(x_42); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } } else { -uint8_t x_54; +lean_object* x_52; +lean_dec(x_21); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_52 = l_Lean_Meta_Grind_markAsInconsistent(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; uint8_t x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = 1; +x_55 = 0; +x_56 = lean_box(0); +x_57 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__6(x_19, x_29, x_3, x_4, x_5, x_6, x_1, x_2, x_54, x_55, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_53); +return x_57; +} +else +{ +uint8_t x_58; +lean_dec(x_29); +lean_dec(x_19); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_52); +if (x_58 == 0) +{ +return x_52; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_52, 0); +x_60 = lean_ctor_get(x_52, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_52); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +} +} +} +} +else +{ +uint8_t x_62; lean_dec(x_21); lean_dec(x_19); lean_dec(x_17); @@ -7494,29 +8115,29 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = !lean_is_exclusive(x_22); -if (x_54 == 0) +x_62 = !lean_is_exclusive(x_22); +if (x_62 == 0) { return x_22; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_22, 0); -x_56 = lean_ctor_get(x_22, 1); -lean_inc(x_56); -lean_inc(x_55); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_22, 0); +x_64 = lean_ctor_get(x_22, 1); +lean_inc(x_64); +lean_inc(x_63); lean_dec(x_22); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } else { -uint8_t x_58; +uint8_t x_66; lean_dec(x_17); lean_dec(x_15); lean_dec(x_14); @@ -7531,23 +8152,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_58 = !lean_is_exclusive(x_18); -if (x_58 == 0) +x_66 = !lean_is_exclusive(x_18); +if (x_66 == 0) { return x_18; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_18, 0); -x_60 = lean_ctor_get(x_18, 1); -lean_inc(x_60); -lean_inc(x_59); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_18, 0); +x_68 = lean_ctor_get(x_18, 1); +lean_inc(x_68); +lean_inc(x_67); lean_dec(x_18); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; } } } @@ -7575,7 +8196,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Gr { lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; x_17 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__8___closed__2; -x_18 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); x_20 = lean_unbox(x_19); @@ -7592,53 +8213,50 @@ return x_23; } else { -if (x_4 == 0) -{ -uint8_t x_24; -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + x_25 = x_18; +} else { + lean_dec_ref(x_18); + x_25 = lean_box(0); +} +x_26 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_18, 1); -x_26 = lean_ctor_get(x_18, 0); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); lean_dec(x_26); +if (x_4 == 0) +{ +lean_object* x_39; lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_6); lean_inc(x_5); -x_27 = l_Lean_Meta_mkEq(x_5, x_6, x_12, x_13, x_14, x_15, x_25); -if (lean_obj_tag(x_27) == 0) +x_39 = l_Lean_Meta_mkEq(x_5, x_6, x_12, x_13, x_14, x_15, x_27); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_MessageData_ofExpr(x_28); -x_31 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_18, 7); -lean_ctor_set(x_18, 1, x_30); -lean_ctor_set(x_18, 0, x_31); -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_18); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_32, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_34, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_35); -lean_dec(x_34); -return x_36; +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_28 = x_40; +x_29 = x_41; +goto block_38; } else { -uint8_t x_37; -lean_free_object(x_18); +uint8_t x_42; +lean_dec(x_25); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -7652,151 +8270,52 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_37 = !lean_is_exclusive(x_27); -if (x_37 == 0) -{ -return x_27; -} -else +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_27, 0); -x_39 = lean_ctor_get(x_27, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_27); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} -} +return x_39; } else { -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_18, 1); -lean_inc(x_41); -lean_dec(x_18); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_6); -lean_inc(x_5); -x_42 = l_Lean_Meta_mkEq(x_5, x_6, x_12, x_13, x_14, x_15, x_41); -if (lean_obj_tag(x_42) == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_39, 0); +x_44 = lean_ctor_get(x_39, 1); lean_inc(x_44); -lean_dec(x_42); -x_45 = l_Lean_MessageData_ofExpr(x_43); -x_46 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_45); -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -x_49 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_48, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_44); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_50, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_51); -lean_dec(x_50); -return x_52; -} -else -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_53 = lean_ctor_get(x_42, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_42, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_55 = x_42; -} else { - lean_dec_ref(x_42); - x_55 = lean_box(0); -} -if (lean_is_scalar(x_55)) { - x_56 = lean_alloc_ctor(1, 2, 0); -} else { - x_56 = x_55; -} -lean_ctor_set(x_56, 0, x_53); -lean_ctor_set(x_56, 1, x_54); -return x_56; +lean_inc(x_43); +lean_dec(x_39); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_18); -if (x_57 == 0) -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_18, 1); -x_59 = lean_ctor_get(x_18, 0); -lean_dec(x_59); +lean_object* x_46; lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_6); lean_inc(x_5); -x_60 = l_Lean_Meta_mkHEq(x_5, x_6, x_12, x_13, x_14, x_15, x_58); -if (lean_obj_tag(x_60) == 0) +x_46 = l_Lean_Meta_mkHEq(x_5, x_6, x_12, x_13, x_14, x_15, x_27); +if (lean_obj_tag(x_46) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -lean_dec(x_60); -x_63 = l_Lean_MessageData_ofExpr(x_61); -x_64 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_18, 7); -lean_ctor_set(x_18, 1, x_63); -lean_ctor_set(x_18, 0, x_64); -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_18); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_62); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -x_69 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_67, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_68); -lean_dec(x_67); -return x_69; +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_28 = x_47; +x_29 = x_48; +goto block_38; } else { -uint8_t x_70; -lean_free_object(x_18); +uint8_t x_49; +lean_dec(x_25); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -7810,68 +8329,57 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_70 = !lean_is_exclusive(x_60); -if (x_70 == 0) +x_49 = !lean_is_exclusive(x_46); +if (x_49 == 0) { -return x_60; -} -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_60, 0); -x_72 = lean_ctor_get(x_60, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_60); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -return x_73; -} -} +return x_46; } else { -lean_object* x_74; lean_object* x_75; -x_74 = lean_ctor_get(x_18, 1); -lean_inc(x_74); -lean_dec(x_18); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_6); -lean_inc(x_5); -x_75 = l_Lean_Meta_mkHEq(x_5, x_6, x_12, x_13, x_14, x_15, x_74); -if (lean_obj_tag(x_75) == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -lean_dec(x_75); -x_78 = l_Lean_MessageData_ofExpr(x_76); -x_79 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_78); -x_81 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_79); -x_82 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_17, x_81, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_77); -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); -lean_inc(x_84); -lean_dec(x_82); -x_85 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_83, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_84); -lean_dec(x_83); -return x_85; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_46, 0); +x_51 = lean_ctor_get(x_46, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_46); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +} +block_38: +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_30 = l_Lean_MessageData_ofExpr(x_28); +x_31 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_25)) { + x_32 = lean_alloc_ctor(7, 2, 0); +} else { + x_32 = x_25; + lean_ctor_set_tag(x_32, 7); +} +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_17, x_33, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_35, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_36); +lean_dec(x_35); +return x_37; +} } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +uint8_t x_53; +lean_dec(x_25); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -7885,27 +8393,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_86 = lean_ctor_get(x_75, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_75, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_88 = x_75; -} else { - lean_dec_ref(x_75); - x_88 = lean_box(0); -} -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); -} else { - x_89 = x_88; -} -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_87); -return x_89; +x_53 = !lean_is_exclusive(x_26); +if (x_53 == 0) +{ +return x_26; } +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_26, 0); +x_55 = lean_ctor_get(x_26, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_26); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } } @@ -8009,7 +8513,7 @@ lean_dec(x_18); lean_dec(x_15); lean_dec(x_3); x_25 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep_go___closed__1; -x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); x_27 = !lean_is_exclusive(x_26); if (x_27 == 0) { @@ -8031,232 +8535,315 @@ return x_33; } else { -lean_object* x_34; uint8_t x_35; -x_34 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +lean_object* x_34; +x_34 = l_Lean_Meta_Grind_updateLastTag(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_35); lean_dec(x_1); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_36 = lean_ctor_get(x_34, 0); -x_37 = lean_ctor_get(x_34, 1); -x_38 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_37); +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_36, 1); +x_40 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_39); lean_dec(x_2); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_40 = lean_ctor_get(x_38, 0); -x_41 = lean_ctor_get(x_38, 1); -x_42 = l_Lean_MessageData_ofFormat(x_36); -x_43 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_42); -lean_ctor_set(x_38, 0, x_43); -x_44 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; -lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_44); -lean_ctor_set(x_34, 0, x_38); -x_45 = l_Lean_MessageData_ofFormat(x_40); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_42 = lean_ctor_get(x_40, 0); +x_43 = lean_ctor_get(x_40, 1); +x_44 = l_Lean_MessageData_ofFormat(x_38); +x_45 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +lean_ctor_set_tag(x_40, 7); +lean_ctor_set(x_40, 1, x_44); +lean_ctor_set(x_40, 0, x_45); +x_46 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +lean_ctor_set_tag(x_36, 7); +lean_ctor_set(x_36, 1, x_46); +lean_ctor_set(x_36, 0, x_40); +x_47 = l_Lean_MessageData_ofFormat(x_42); lean_ctor_set_tag(x_26, 7); -lean_ctor_set(x_26, 1, x_45); -lean_ctor_set(x_26, 0, x_34); -x_46 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_26); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_41); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_apply_10(x_30, x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_50); -return x_51; +lean_ctor_set(x_26, 1, x_47); +lean_ctor_set(x_26, 0, x_36); +x_48 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_26); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_43); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_apply_10(x_30, x_51, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_52); +return x_53; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_52 = lean_ctor_get(x_38, 0); -x_53 = lean_ctor_get(x_38, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_38); -x_54 = l_Lean_MessageData_ofFormat(x_36); -x_55 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -x_57 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; -lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_57); -lean_ctor_set(x_34, 0, x_56); -x_58 = l_Lean_MessageData_ofFormat(x_52); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_54 = lean_ctor_get(x_40, 0); +x_55 = lean_ctor_get(x_40, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_40); +x_56 = l_Lean_MessageData_ofFormat(x_38); +x_57 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_56); +x_59 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +lean_ctor_set_tag(x_36, 7); +lean_ctor_set(x_36, 1, x_59); +lean_ctor_set(x_36, 0, x_58); +x_60 = l_Lean_MessageData_ofFormat(x_54); lean_ctor_set_tag(x_26, 7); -lean_ctor_set(x_26, 1, x_58); -lean_ctor_set(x_26, 0, x_34); -x_59 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -x_60 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_60, 0, x_26); -lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_60, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_53); -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_apply_10(x_30, x_62, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_63); -return x_64; +lean_ctor_set(x_26, 1, x_60); +lean_ctor_set(x_26, 0, x_36); +x_61 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_62 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_62, 0, x_26); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_62, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_55); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_apply_10(x_30, x_64, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_65); +return x_66; } } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_65 = lean_ctor_get(x_34, 0); -x_66 = lean_ctor_get(x_34, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_34); -x_67 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_66); -lean_dec(x_2); -x_68 = lean_ctor_get(x_67, 0); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_67 = lean_ctor_get(x_36, 0); +x_68 = lean_ctor_get(x_36, 1); lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_70 = x_67; +lean_inc(x_67); +lean_dec(x_36); +x_69 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_68); +lean_dec(x_2); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_72 = x_69; } else { - lean_dec_ref(x_67); - x_70 = lean_box(0); + lean_dec_ref(x_69); + x_72 = lean_box(0); } -x_71 = l_Lean_MessageData_ofFormat(x_65); -x_72 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_70)) { - x_73 = lean_alloc_ctor(7, 2, 0); +x_73 = l_Lean_MessageData_ofFormat(x_67); +x_74 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_72)) { + x_75 = lean_alloc_ctor(7, 2, 0); } else { - x_73 = x_70; - lean_ctor_set_tag(x_73, 7); + x_75 = x_72; + lean_ctor_set_tag(x_75, 7); } -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_71); -x_74 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Lean_MessageData_ofFormat(x_68); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_73); +x_76 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +x_78 = l_Lean_MessageData_ofFormat(x_70); lean_ctor_set_tag(x_26, 7); -lean_ctor_set(x_26, 1, x_76); -lean_ctor_set(x_26, 0, x_75); -x_77 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_26); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_78, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_69); -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -lean_dec(x_79); -x_82 = lean_apply_10(x_30, x_80, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_81); -return x_82; +lean_ctor_set(x_26, 1, x_78); +lean_ctor_set(x_26, 0, x_77); +x_79 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_26); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_80, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_71); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_apply_10(x_30, x_82, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_83); +return x_84; +} +} +else +{ +uint8_t x_85; +lean_free_object(x_26); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_85 = !lean_is_exclusive(x_34); +if (x_85 == 0) +{ +return x_34; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_34, 0); +x_87 = lean_ctor_get(x_34, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_34); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; +} } } } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_83 = lean_ctor_get(x_26, 0); -x_84 = lean_ctor_get(x_26, 1); -lean_inc(x_84); -lean_inc(x_83); +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_89 = lean_ctor_get(x_26, 0); +x_90 = lean_ctor_get(x_26, 1); +lean_inc(x_90); +lean_inc(x_89); lean_dec(x_26); -x_85 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1; -x_86 = lean_unbox(x_83); -lean_dec(x_83); -if (x_86 == 0) +x_91 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__1; +x_92 = lean_unbox(x_89); +lean_dec(x_89); +if (x_92 == 0) { -lean_object* x_87; lean_object* x_88; +lean_object* x_93; lean_object* x_94; lean_dec(x_2); lean_dec(x_1); -x_87 = lean_box(0); -x_88 = lean_apply_10(x_85, x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_84); -return x_88; +x_93 = lean_box(0); +x_94 = lean_apply_10(x_91, x_93, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_90); +return x_94; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_89 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_84); +lean_object* x_95; +x_95 = l_Lean_Meta_Grind_updateLastTag(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_90); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_97 = l_Lean_Meta_Grind_ppENodeRef(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_96); lean_dec(x_1); -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_89, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_92 = x_89; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_100 = x_97; } else { - lean_dec_ref(x_89); - x_92 = lean_box(0); + lean_dec_ref(x_97); + x_100 = lean_box(0); } -x_93 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_91); +x_101 = l_Lean_Meta_Grind_ppENodeRef(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_99); lean_dec(x_2); -x_94 = lean_ctor_get(x_93, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_96 = x_93; +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_101, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_104 = x_101; } else { - lean_dec_ref(x_93); - x_96 = lean_box(0); + lean_dec_ref(x_101); + x_104 = lean_box(0); } -x_97 = l_Lean_MessageData_ofFormat(x_90); -x_98 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -if (lean_is_scalar(x_96)) { - x_99 = lean_alloc_ctor(7, 2, 0); +x_105 = l_Lean_MessageData_ofFormat(x_98); +x_106 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +if (lean_is_scalar(x_104)) { + x_107 = lean_alloc_ctor(7, 2, 0); } else { - x_99 = x_96; - lean_ctor_set_tag(x_99, 7); + x_107 = x_104; + lean_ctor_set_tag(x_107, 7); } -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set(x_99, 1, x_97); -x_100 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; -if (lean_is_scalar(x_92)) { - x_101 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_105); +x_108 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__3; +if (lean_is_scalar(x_100)) { + x_109 = lean_alloc_ctor(7, 2, 0); } else { - x_101 = x_92; - lean_ctor_set_tag(x_101, 7); + x_109 = x_100; + lean_ctor_set_tag(x_109, 7); +} +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_MessageData_ofFormat(x_102); +x_111 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +x_112 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_113, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_103); +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_114, 1); +lean_inc(x_116); +lean_dec(x_114); +x_117 = lean_apply_10(x_91, x_115, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_116); +return x_117; +} +else +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_118 = lean_ctor_get(x_95, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_95, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_120 = x_95; +} else { + lean_dec_ref(x_95); + x_120 = lean_box(0); +} +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); +} else { + x_121 = x_120; +} +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; } -lean_ctor_set(x_101, 0, x_99); -lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean_MessageData_ofFormat(x_94); -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -x_104 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqStep___closed__5; -x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_25, x_105, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_95); -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); -lean_dec(x_106); -x_109 = lean_apply_10(x_85, x_107, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_108); -return x_109; } } } } else { -uint8_t x_110; +uint8_t x_122; lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); @@ -8269,29 +8856,29 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_110 = !lean_is_exclusive(x_17); -if (x_110 == 0) +x_122 = !lean_is_exclusive(x_17); +if (x_122 == 0) { return x_17; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_17, 0); -x_112 = lean_ctor_get(x_17, 1); -lean_inc(x_112); -lean_inc(x_111); +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_17, 0); +x_124 = lean_ctor_get(x_17, 1); +lean_inc(x_124); +lean_inc(x_123); lean_dec(x_17); -x_113 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -return x_113; +x_125 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; } } } else { -uint8_t x_114; +uint8_t x_126; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -8303,23 +8890,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_114 = !lean_is_exclusive(x_14); -if (x_114 == 0) +x_126 = !lean_is_exclusive(x_14); +if (x_126 == 0) { return x_14; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_14, 0); -x_116 = lean_ctor_get(x_14, 1); -lean_inc(x_116); -lean_inc(x_115); +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_14, 0); +x_128 = lean_ctor_get(x_14, 1); +lean_inc(x_128); +lean_inc(x_127); lean_dec(x_14); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +x_129 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +return x_129; } } } @@ -8514,13 +9101,13 @@ return x_22; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; x_23 = lean_ctor_get(x_11, 0); x_24 = lean_ctor_get(x_11, 1); x_25 = lean_ctor_get(x_11, 2); x_26 = lean_ctor_get(x_11, 3); x_27 = lean_ctor_get(x_11, 4); -x_28 = lean_ctor_get_uint8(x_11, sizeof(void*)*14); +x_28 = lean_ctor_get_uint8(x_11, sizeof(void*)*19); x_29 = lean_ctor_get(x_11, 6); x_30 = lean_ctor_get(x_11, 7); x_31 = lean_ctor_get(x_11, 8); @@ -8529,6 +9116,16 @@ x_33 = lean_ctor_get(x_11, 10); x_34 = lean_ctor_get(x_11, 11); x_35 = lean_ctor_get(x_11, 12); x_36 = lean_ctor_get(x_11, 13); +x_37 = lean_ctor_get(x_11, 14); +x_38 = lean_ctor_get(x_11, 15); +x_39 = lean_ctor_get(x_11, 16); +x_40 = lean_ctor_get(x_11, 17); +x_41 = lean_ctor_get(x_11, 18); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); @@ -8543,43 +9140,48 @@ lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_dec(x_11); -x_37 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1; -x_38 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_38, 0, x_23); -lean_ctor_set(x_38, 1, x_24); -lean_ctor_set(x_38, 2, x_25); -lean_ctor_set(x_38, 3, x_26); -lean_ctor_set(x_38, 4, x_27); -lean_ctor_set(x_38, 5, x_37); -lean_ctor_set(x_38, 6, x_29); -lean_ctor_set(x_38, 7, x_30); -lean_ctor_set(x_38, 8, x_31); -lean_ctor_set(x_38, 9, x_32); -lean_ctor_set(x_38, 10, x_33); -lean_ctor_set(x_38, 11, x_34); -lean_ctor_set(x_38, 12, x_35); -lean_ctor_set(x_38, 13, x_36); -lean_ctor_set_uint8(x_38, sizeof(void*)*14, x_28); -x_39 = lean_st_ref_set(x_1, x_38, x_12); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; +x_42 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_resetNewEqs___closed__1; +x_43 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_43, 0, x_23); +lean_ctor_set(x_43, 1, x_24); +lean_ctor_set(x_43, 2, x_25); +lean_ctor_set(x_43, 3, x_26); +lean_ctor_set(x_43, 4, x_27); +lean_ctor_set(x_43, 5, x_42); +lean_ctor_set(x_43, 6, x_29); +lean_ctor_set(x_43, 7, x_30); +lean_ctor_set(x_43, 8, x_31); +lean_ctor_set(x_43, 9, x_32); +lean_ctor_set(x_43, 10, x_33); +lean_ctor_set(x_43, 11, x_34); +lean_ctor_set(x_43, 12, x_35); +lean_ctor_set(x_43, 13, x_36); +lean_ctor_set(x_43, 14, x_37); +lean_ctor_set(x_43, 15, x_38); +lean_ctor_set(x_43, 16, x_39); +lean_ctor_set(x_43, 17, x_40); +lean_ctor_set(x_43, 18, x_41); +lean_ctor_set_uint8(x_43, sizeof(void*)*19, x_28); +x_44 = lean_st_ref_set(x_1, x_43, x_12); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_46 = x_44; } else { - lean_dec_ref(x_39); - x_41 = lean_box(0); + lean_dec_ref(x_44); + x_46 = lean_box(0); } -x_42 = lean_box(0); -if (lean_is_scalar(x_41)) { - x_43 = lean_alloc_ctor(0, 2, 0); +x_47 = lean_box(0); +if (lean_is_scalar(x_46)) { + x_48 = lean_alloc_ctor(0, 2, 0); } else { - x_43 = x_41; + x_48 = x_46; } -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_45); +return x_48; } } } @@ -8671,14 +9273,14 @@ return x_26; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; x_27 = lean_ctor_get(x_17, 0); x_28 = lean_ctor_get(x_17, 1); x_29 = lean_ctor_get(x_17, 2); x_30 = lean_ctor_get(x_17, 3); x_31 = lean_ctor_get(x_17, 4); x_32 = lean_ctor_get(x_17, 5); -x_33 = lean_ctor_get_uint8(x_17, sizeof(void*)*14); +x_33 = lean_ctor_get_uint8(x_17, sizeof(void*)*19); x_34 = lean_ctor_get(x_17, 6); x_35 = lean_ctor_get(x_17, 7); x_36 = lean_ctor_get(x_17, 8); @@ -8687,6 +9289,16 @@ x_38 = lean_ctor_get(x_17, 10); x_39 = lean_ctor_get(x_17, 11); x_40 = lean_ctor_get(x_17, 12); x_41 = lean_ctor_get(x_17, 13); +x_42 = lean_ctor_get(x_17, 14); +x_43 = lean_ctor_get(x_17, 15); +x_44 = lean_ctor_get(x_17, 16); +x_45 = lean_ctor_get(x_17, 17); +x_46 = lean_ctor_get(x_17, 18); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); lean_inc(x_41); lean_inc(x_40); lean_inc(x_39); @@ -8699,167 +9311,192 @@ lean_inc(x_32); lean_inc(x_31); lean_inc(x_30); lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_17); -x_42 = lean_array_pop(x_32); -x_43 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_43, 0, x_27); -lean_ctor_set(x_43, 1, x_28); -lean_ctor_set(x_43, 2, x_29); -lean_ctor_set(x_43, 3, x_30); -lean_ctor_set(x_43, 4, x_31); -lean_ctor_set(x_43, 5, x_42); -lean_ctor_set(x_43, 6, x_34); -lean_ctor_set(x_43, 7, x_35); -lean_ctor_set(x_43, 8, x_36); -lean_ctor_set(x_43, 9, x_37); -lean_ctor_set(x_43, 10, x_38); -lean_ctor_set(x_43, 11, x_39); -lean_ctor_set(x_43, 12, x_40); -lean_ctor_set(x_43, 13, x_41); -lean_ctor_set_uint8(x_43, sizeof(void*)*14, x_33); -x_44 = lean_st_ref_set(x_1, x_43, x_18); -x_45 = lean_ctor_get(x_44, 1); -lean_inc(x_45); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - lean_ctor_release(x_44, 1); - x_46 = x_44; +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_17); +x_47 = lean_array_pop(x_32); +x_48 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_48, 0, x_27); +lean_ctor_set(x_48, 1, x_28); +lean_ctor_set(x_48, 2, x_29); +lean_ctor_set(x_48, 3, x_30); +lean_ctor_set(x_48, 4, x_31); +lean_ctor_set(x_48, 5, x_47); +lean_ctor_set(x_48, 6, x_34); +lean_ctor_set(x_48, 7, x_35); +lean_ctor_set(x_48, 8, x_36); +lean_ctor_set(x_48, 9, x_37); +lean_ctor_set(x_48, 10, x_38); +lean_ctor_set(x_48, 11, x_39); +lean_ctor_set(x_48, 12, x_40); +lean_ctor_set(x_48, 13, x_41); +lean_ctor_set(x_48, 14, x_42); +lean_ctor_set(x_48, 15, x_43); +lean_ctor_set(x_48, 16, x_44); +lean_ctor_set(x_48, 17, x_45); +lean_ctor_set(x_48, 18, x_46); +lean_ctor_set_uint8(x_48, sizeof(void*)*19, x_33); +x_49 = lean_st_ref_set(x_1, x_48, x_18); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_51 = x_49; } else { - lean_dec_ref(x_44); - x_46 = lean_box(0); + lean_dec_ref(x_49); + x_51 = lean_box(0); } -if (lean_is_scalar(x_46)) { - x_47 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(0, 2, 0); } else { - x_47 = x_46; + x_52 = x_51; } -lean_ctor_set(x_47, 0, x_15); -lean_ctor_set(x_47, 1, x_45); -return x_47; +lean_ctor_set(x_52, 0, x_15); +lean_ctor_set(x_52, 1, x_50); +return x_52; } } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_48 = lean_ctor_get(x_10, 0); -x_49 = lean_ctor_get(x_10, 1); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = lean_ctor_get(x_10, 0); +x_54 = lean_ctor_get(x_10, 1); +lean_inc(x_54); +lean_inc(x_53); lean_dec(x_10); -x_50 = lean_ctor_get(x_48, 5); -lean_inc(x_50); -lean_dec(x_48); -x_51 = l_Array_back_x3f___rarg(x_50); -lean_dec(x_50); -if (lean_obj_tag(x_51) == 0) +x_55 = lean_ctor_get(x_53, 5); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Array_back_x3f___rarg(x_55); +lean_dec(x_55); +if (lean_obj_tag(x_56) == 0) { -lean_object* x_52; -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_49); -return x_52; +lean_object* x_57; +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +return x_57; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_53 = lean_st_ref_take(x_1, x_49); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_ctor_get(x_54, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -x_58 = lean_ctor_get(x_54, 2); -lean_inc(x_58); -x_59 = lean_ctor_get(x_54, 3); +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_58 = lean_st_ref_take(x_1, x_54); +x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); -x_60 = lean_ctor_get(x_54, 4); +x_60 = lean_ctor_get(x_58, 1); lean_inc(x_60); -x_61 = lean_ctor_get(x_54, 5); +lean_dec(x_58); +x_61 = lean_ctor_get(x_59, 0); lean_inc(x_61); -x_62 = lean_ctor_get_uint8(x_54, sizeof(void*)*14); -x_63 = lean_ctor_get(x_54, 6); +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +x_63 = lean_ctor_get(x_59, 2); lean_inc(x_63); -x_64 = lean_ctor_get(x_54, 7); +x_64 = lean_ctor_get(x_59, 3); lean_inc(x_64); -x_65 = lean_ctor_get(x_54, 8); +x_65 = lean_ctor_get(x_59, 4); lean_inc(x_65); -x_66 = lean_ctor_get(x_54, 9); +x_66 = lean_ctor_get(x_59, 5); lean_inc(x_66); -x_67 = lean_ctor_get(x_54, 10); -lean_inc(x_67); -x_68 = lean_ctor_get(x_54, 11); +x_67 = lean_ctor_get_uint8(x_59, sizeof(void*)*19); +x_68 = lean_ctor_get(x_59, 6); lean_inc(x_68); -x_69 = lean_ctor_get(x_54, 12); +x_69 = lean_ctor_get(x_59, 7); lean_inc(x_69); -x_70 = lean_ctor_get(x_54, 13); +x_70 = lean_ctor_get(x_59, 8); lean_inc(x_70); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - lean_ctor_release(x_54, 2); - lean_ctor_release(x_54, 3); - lean_ctor_release(x_54, 4); - lean_ctor_release(x_54, 5); - lean_ctor_release(x_54, 6); - lean_ctor_release(x_54, 7); - lean_ctor_release(x_54, 8); - lean_ctor_release(x_54, 9); - lean_ctor_release(x_54, 10); - lean_ctor_release(x_54, 11); - lean_ctor_release(x_54, 12); - lean_ctor_release(x_54, 13); - x_71 = x_54; +x_71 = lean_ctor_get(x_59, 9); +lean_inc(x_71); +x_72 = lean_ctor_get(x_59, 10); +lean_inc(x_72); +x_73 = lean_ctor_get(x_59, 11); +lean_inc(x_73); +x_74 = lean_ctor_get(x_59, 12); +lean_inc(x_74); +x_75 = lean_ctor_get(x_59, 13); +lean_inc(x_75); +x_76 = lean_ctor_get(x_59, 14); +lean_inc(x_76); +x_77 = lean_ctor_get(x_59, 15); +lean_inc(x_77); +x_78 = lean_ctor_get(x_59, 16); +lean_inc(x_78); +x_79 = lean_ctor_get(x_59, 17); +lean_inc(x_79); +x_80 = lean_ctor_get(x_59, 18); +lean_inc(x_80); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + lean_ctor_release(x_59, 2); + lean_ctor_release(x_59, 3); + lean_ctor_release(x_59, 4); + lean_ctor_release(x_59, 5); + lean_ctor_release(x_59, 6); + lean_ctor_release(x_59, 7); + lean_ctor_release(x_59, 8); + lean_ctor_release(x_59, 9); + lean_ctor_release(x_59, 10); + lean_ctor_release(x_59, 11); + lean_ctor_release(x_59, 12); + lean_ctor_release(x_59, 13); + lean_ctor_release(x_59, 14); + lean_ctor_release(x_59, 15); + lean_ctor_release(x_59, 16); + lean_ctor_release(x_59, 17); + lean_ctor_release(x_59, 18); + x_81 = x_59; } else { - lean_dec_ref(x_54); - x_71 = lean_box(0); + lean_dec_ref(x_59); + x_81 = lean_box(0); } -x_72 = lean_array_pop(x_61); -if (lean_is_scalar(x_71)) { - x_73 = lean_alloc_ctor(0, 14, 1); +x_82 = lean_array_pop(x_66); +if (lean_is_scalar(x_81)) { + x_83 = lean_alloc_ctor(0, 19, 1); } else { - x_73 = x_71; -} -lean_ctor_set(x_73, 0, x_56); -lean_ctor_set(x_73, 1, x_57); -lean_ctor_set(x_73, 2, x_58); -lean_ctor_set(x_73, 3, x_59); -lean_ctor_set(x_73, 4, x_60); -lean_ctor_set(x_73, 5, x_72); -lean_ctor_set(x_73, 6, x_63); -lean_ctor_set(x_73, 7, x_64); -lean_ctor_set(x_73, 8, x_65); -lean_ctor_set(x_73, 9, x_66); -lean_ctor_set(x_73, 10, x_67); -lean_ctor_set(x_73, 11, x_68); -lean_ctor_set(x_73, 12, x_69); -lean_ctor_set(x_73, 13, x_70); -lean_ctor_set_uint8(x_73, sizeof(void*)*14, x_62); -x_74 = lean_st_ref_set(x_1, x_73, x_55); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; + x_83 = x_81; +} +lean_ctor_set(x_83, 0, x_61); +lean_ctor_set(x_83, 1, x_62); +lean_ctor_set(x_83, 2, x_63); +lean_ctor_set(x_83, 3, x_64); +lean_ctor_set(x_83, 4, x_65); +lean_ctor_set(x_83, 5, x_82); +lean_ctor_set(x_83, 6, x_68); +lean_ctor_set(x_83, 7, x_69); +lean_ctor_set(x_83, 8, x_70); +lean_ctor_set(x_83, 9, x_71); +lean_ctor_set(x_83, 10, x_72); +lean_ctor_set(x_83, 11, x_73); +lean_ctor_set(x_83, 12, x_74); +lean_ctor_set(x_83, 13, x_75); +lean_ctor_set(x_83, 14, x_76); +lean_ctor_set(x_83, 15, x_77); +lean_ctor_set(x_83, 16, x_78); +lean_ctor_set(x_83, 17, x_79); +lean_ctor_set(x_83, 18, x_80); +lean_ctor_set_uint8(x_83, sizeof(void*)*19, x_67); +x_84 = lean_st_ref_set(x_1, x_83, x_60); +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_86 = x_84; } else { - lean_dec_ref(x_74); - x_76 = lean_box(0); + lean_dec_ref(x_84); + x_86 = lean_box(0); } -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(0, 2, 0); } else { - x_77 = x_76; + x_87 = x_86; } -lean_ctor_set(x_77, 0, x_51); -lean_ctor_set(x_77, 1, x_75); -return x_77; +lean_ctor_set(x_87, 0, x_56); +lean_ctor_set(x_87, 1, x_85); +return x_87; } } } @@ -9204,6 +9841,117 @@ x_14 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore(x_1, return x_14; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_14 = l_Lean_Meta_Grind_internalize(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_16 = l_Lean_Meta_Grind_internalize(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = 0; +x_19 = l___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_addEqCore(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_16); +if (x_20 == 0) +{ +return x_16; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_16, 0); +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_16); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +else +{ +uint8_t x_24; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_14); +if (x_24 == 0) +{ +return x_14; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_14, 0); +x_26 = lean_ctor_get(x_14, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_14); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add_goEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { @@ -9978,7 +10726,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_add(lean_object* x_1, lean_object* x_ { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = l_Lean_Meta_Grind_add___closed__2; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_unbox(x_15); @@ -9999,53 +10747,138 @@ uint8_t x_20; x_20 = !lean_is_exclusive(x_14); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_14, 1); x_22 = lean_ctor_get(x_14, 0); lean_dec(x_22); +x_23 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); lean_inc(x_1); -x_23 = l_Lean_MessageData_ofExpr(x_1); -x_24 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_25 = l_Lean_MessageData_ofExpr(x_1); +x_26 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; lean_ctor_set_tag(x_14, 7); -lean_ctor_set(x_14, 1, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Meta_Grind_add___lambda__4(x_1, x_2, x_3, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); -lean_dec(x_27); -return x_29; +lean_ctor_set(x_14, 1, x_25); +lean_ctor_set(x_14, 0, x_26); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Meta_Grind_add___lambda__4(x_1, x_2, x_3, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_29); +return x_31; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); +uint8_t x_32; +lean_free_object(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) +{ +return x_23; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_23, 0); +x_34 = lean_ctor_get(x_23, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_23); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); lean_dec(x_14); +x_37 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_36); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); lean_inc(x_1); -x_31 = l_Lean_MessageData_ofExpr(x_1); -x_32 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_Meta_Grind_add___lambda__4(x_1, x_2, x_3, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); -lean_dec(x_36); -return x_38; +x_39 = l_Lean_MessageData_ofExpr(x_1); +x_40 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Core_0__Lean_Meta_Grind_removeParents___spec__4___closed__8; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_add___lambda__4(x_1, x_2, x_3, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); +lean_dec(x_44); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_47 = lean_ctor_get(x_37, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_37, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_49 = x_37; +} else { + lean_dec_ref(x_37); + x_49 = lean_box(0); +} +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(1, 2, 0); +} else { + x_50 = x_49; +} +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c index 45be339f7150..975d97d21fe8 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Ctor.c @@ -20,20 +20,17 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_p lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_lor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); -lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__3; -static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; static lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__3___closed__2; lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__5; lean_object* l_Lean_stringToMessageData(lean_object*); @@ -42,7 +39,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__3___boxed(lea LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getFalseExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__1; -lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); @@ -53,27 +49,26 @@ lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7; lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_Grind_propagateCtor___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_Grind_propagateCtor___spec__1___closed__1; -static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__1; -double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_propagateCtor___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_grind_mk_eq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); @@ -88,12 +83,10 @@ uint64_t lean_uint64_shift_left(uint64_t, uint64_t); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_Grind_propagateCtor___spec__1___closed__3; lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1; -lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_Grind_propagateCtor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__3___closed__1; -lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateCtor___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_Grind_propagateCtor___spec__1___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__2; @@ -102,394 +95,46 @@ uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__2; static uint64_t l_Lean_Meta_Grind_propagateCtor___closed__1; -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 2); -x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -} -static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; double x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = 0; -x_3 = l_Float_ofScientific(x_1, x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_12 = lean_ctor_get(x_9, 5); -x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_13 = l_Lean_Meta_Grind_shareCommon(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); -x_16 = lean_st_ref_take(x_10, x_15); +x_16 = l_Lean_Meta_Grind_shareCommon(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 3); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 0); -lean_dec(x_21); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_17, 3); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) -{ -lean_object* x_25; double x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_25 = lean_ctor_get(x_18, 0); -x_26 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; -x_27 = 0; -x_28 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; -x_29 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); -lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); -x_30 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; -x_31 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_14); -lean_ctor_set(x_31, 2, x_30); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_31); -lean_ctor_set(x_16, 0, x_12); -x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); -lean_ctor_set(x_18, 0, x_32); -x_33 = lean_st_ref_set(x_10, x_17, x_20); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} -} -else -{ -uint64_t x_40; lean_object* x_41; double x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_41 = lean_ctor_get(x_18, 0); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; -x_43 = 0; -x_44 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; -x_45 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); -lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); -x_46 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; -x_47 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_46); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_16, 0, x_12); -x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); -x_49 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); -lean_ctor_set(x_17, 3, x_49); -x_50 = lean_st_ref_set(x_10, x_17, x_20); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_55 = lean_ctor_get(x_17, 0); -x_56 = lean_ctor_get(x_17, 1); -x_57 = lean_ctor_get(x_17, 2); -x_58 = lean_ctor_get(x_17, 4); -x_59 = lean_ctor_get(x_17, 5); -x_60 = lean_ctor_get(x_17, 6); -x_61 = lean_ctor_get(x_17, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_17); -x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_63 = lean_ctor_get(x_18, 0); -lean_inc(x_63); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_64 = x_18; -} else { - lean_dec_ref(x_18); - x_64 = lean_box(0); -} -x_65 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; -x_66 = 0; -x_67 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; -x_68 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); -lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); -x_69 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; -x_70 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_14); -lean_ctor_set(x_70, 2, x_69); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_70); -lean_ctor_set(x_16, 0, x_12); -x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); -if (lean_is_scalar(x_64)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_64; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); -x_73 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_73, 0, x_55); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_57); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_73, 4, x_58); -lean_ctor_set(x_73, 5, x_59); -lean_ctor_set(x_73, 6, x_60); -lean_ctor_set(x_73, 7, x_61); -x_74 = lean_st_ref_set(x_10, x_73, x_20); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); -} -x_77 = lean_box(0); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_75); -return x_78; -} -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_79 = lean_ctor_get(x_16, 1); -lean_inc(x_79); lean_dec(x_16); -x_80 = lean_ctor_get(x_17, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_17, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_17, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_17, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_17, 5); -lean_inc(x_84); -x_85 = lean_ctor_get(x_17, 6); -lean_inc(x_85); -x_86 = lean_ctor_get(x_17, 7); -lean_inc(x_86); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - lean_ctor_release(x_17, 2); - lean_ctor_release(x_17, 3); - lean_ctor_release(x_17, 4); - lean_ctor_release(x_17, 5); - lean_ctor_release(x_17, 6); - lean_ctor_release(x_17, 7); - x_87 = x_17; -} else { - lean_dec_ref(x_17); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_89 = lean_ctor_get(x_18, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_90 = x_18; -} else { - lean_dec_ref(x_18); - x_90 = lean_box(0); -} -x_91 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1; -x_92 = 0; -x_93 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; -x_94 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); -lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); -x_95 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3; -x_96 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_14); -lean_ctor_set(x_96, 2, x_95); -lean_inc(x_12); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_12); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); -if (lean_is_scalar(x_90)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_90; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); -if (lean_is_scalar(x_87)) { - x_100 = lean_alloc_ctor(0, 8, 0); -} else { - x_100 = x_87; -} -lean_ctor_set(x_100, 0, x_80); -lean_ctor_set(x_100, 1, x_81); -lean_ctor_set(x_100, 2, x_82); -lean_ctor_set(x_100, 3, x_99); -lean_ctor_set(x_100, 4, x_83); -lean_ctor_set(x_100, 5, x_84); -lean_ctor_set(x_100, 6, x_85); -lean_ctor_set(x_100, 7, x_86); -x_101 = lean_st_ref_set(x_10, x_100, x_79); -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_103 = x_101; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = 1; -x_14 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_1, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +x_19 = 1; +x_20 = l_Lean_Meta_Grind_pushEqCore(x_14, x_17, x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_20; } } LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_13; lean_object* x_14; -x_13 = 0; -x_14 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_1, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_13 = l_Lean_Meta_Grind_shareCommon(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_shareCommon(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 0; +x_20 = l_Lean_Meta_Grind_pushEqCore(x_14, x_17, x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_20; } } static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1() { @@ -637,8 +282,16 @@ return x_2; static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -648,7 +301,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Gr { lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__3; -x_13 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { @@ -669,65 +322,146 @@ return x_20; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_21 = l_Lean_indentExpr(x_1); -x_22 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; +lean_object* x_21; +x_21 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = l_Lean_indentExpr(x_1); +x_24 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; lean_ctor_set_tag(x_13, 7); -lean_ctor_set(x_13, 1, x_21); -lean_ctor_set(x_13, 0, x_22); -x_23 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7; -x_24 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_24, 0, x_13); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2(x_12, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_apply_10(x_17, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -return x_28; +lean_ctor_set(x_13, 1, x_23); +lean_ctor_set(x_13, 0, x_24); +x_25 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_13); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_12, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_apply_10(x_17, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_29); +return x_30; } +else +{ +uint8_t x_31; +lean_free_object(x_13); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_21); +if (x_31 == 0) +{ +return x_21; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_13, 0); -x_30 = lean_ctor_get(x_13, 1); -lean_inc(x_30); -lean_inc(x_29); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_21); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = lean_ctor_get(x_13, 0); +x_36 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_inc(x_35); lean_dec(x_13); -x_31 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__4; -x_32 = lean_unbox(x_29); -lean_dec(x_29); -if (x_32 == 0) +x_37 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__4; +x_38 = lean_unbox(x_35); +lean_dec(x_35); +if (x_38 == 0) { -lean_object* x_33; lean_object* x_34; +lean_object* x_39; lean_object* x_40; lean_dec(x_1); -x_33 = lean_box(0); -x_34 = lean_apply_10(x_31, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -return x_34; +x_39 = lean_box(0); +x_40 = lean_apply_10(x_37, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +return x_40; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_35 = l_Lean_indentExpr(x_1); -x_36 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_35); -x_38 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2(x_12, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); +lean_object* x_41; +x_41 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_42 = lean_ctor_get(x_41, 1); lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_apply_10(x_31, x_41, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_42); -return x_43; +lean_dec(x_41); +x_43 = l_Lean_indentExpr(x_1); +x_44 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +x_46 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_12, x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_42); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = lean_apply_10(x_37, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_50); +return x_51; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_52 = lean_ctor_get(x_41, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_41, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_54 = x_41; +} else { + lean_dec_ref(x_41); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(1, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} } } } @@ -870,10 +604,9 @@ return x_42; } else { -uint8_t x_43; lean_object* x_44; +lean_object* x_43; lean_dec(x_1); -x_43 = 1; -x_44 = l_Lean_Meta_Grind_pushEqCore(x_31, x_19, x_2, x_43, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_43 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1(x_2, x_31, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -882,18 +615,17 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_44; +return x_43; } } } else { -uint8_t x_45; lean_object* x_46; +lean_object* x_44; lean_dec(x_32); lean_dec(x_31); lean_dec(x_1); -x_45 = 0; -x_46 = l_Lean_Meta_Grind_pushEqCore(x_24, x_19, x_2, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_44 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__2(x_2, x_24, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -902,54 +634,22 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_46; +return x_44; } } } else { -lean_object* x_47; +lean_object* x_45; lean_dec(x_25); lean_dec(x_1); -x_47 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3(x_2, x_24, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_47; +x_45 = l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3(x_2, x_24, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_45; } } } } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -1094,8 +794,7 @@ x_14 = lean_ctor_get(x_11, 1); x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); lean_dec(x_13); -lean_inc(x_1); -x_16 = lean_environment_find(x_15, x_1); +x_16 = l_Lean_Environment_find_x3f(x_15, x_1); if (lean_obj_tag(x_16) == 0) { uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; @@ -1135,8 +834,7 @@ lean_dec(x_11); x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); -lean_inc(x_1); -x_28 = lean_environment_find(x_27, x_1); +x_28 = l_Lean_Environment_find_x3f(x_27, x_1); if (lean_obj_tag(x_28) == 0) { uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -1209,10 +907,6 @@ x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); x_22 = l_Lean_Meta_Grind_closeGoal(x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); return x_22; } else @@ -2294,11 +1988,6 @@ _G_initialized = true; res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__1(); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__2); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___spec__2___closed__3); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__1); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__3___closed__2(); @@ -2317,6 +2006,8 @@ l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lam lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__6); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__7); +l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___lambda__5___closed__8); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__1); l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Ctor_0__Lean_Meta_Grind_propagateInjEqs___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/DoNotSimp.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/DoNotSimp.c new file mode 100644 index 000000000000..2e56424bb556 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/DoNotSimp.c @@ -0,0 +1,608 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.DoNotSimp +// Imports: Init.Grind.Util Init.Simproc Lean.Meta.Tactic.Simp.Simproc +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2; +uint8_t l_Lean_Expr_isApp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +static lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addDoNotSimp___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___closed__2; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_Simprocs_add(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4; +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7; +lean_object* l_Lean_Meta_Simp_registerBuiltinDSimproc(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsDoNotSimp___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9; +static lean_object* l_Lean_Meta_Grind_markAsDoNotSimp___closed__2; +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___closed__1; +static lean_object* l_Lean_Meta_Grind_markAsDoNotSimp___closed__3; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doNotSimp", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__1; +x_2 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__2; +x_3 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsDoNotSimp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +x_9 = lean_array_mk(x_8); +x_10 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +x_11 = l_Lean_Meta_mkAppM(x_10, x_9, x_2, x_3, x_4, x_5, x_6); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +} +static lean_object* _init_l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +} +static lean_object* _init_l_Lean_Meta_Grind_reduceDoNotSimp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___boxed), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_inc(x_1); +x_10 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_Grind_reduceDoNotSimp___closed__1; +x_14 = l_Lean_Expr_cleanupAnnotations(x_11); +x_15 = l_Lean_Expr_isApp(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_14); +lean_dec(x_1); +x_16 = lean_box(0); +x_17 = lean_apply_9(x_13, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_17; +} +else +{ +lean_object* x_18; uint8_t x_19; +x_18 = l_Lean_Expr_appFnCleanup(x_14, lean_box(0)); +x_19 = l_Lean_Expr_isApp(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_18); +lean_dec(x_1); +x_20 = lean_box(0); +x_21 = lean_apply_9(x_13, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = l_Lean_Expr_appFnCleanup(x_18, lean_box(0)); +x_23 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +x_24 = l_Lean_Expr_isConstOf(x_22, x_23); +lean_dec(x_22); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_1); +x_25 = lean_box(0); +x_26 = lean_apply_9(x_13, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_box(0); +x_28 = l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1(x_1, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_28; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_reduceDoNotSimp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("reduceDoNotSimp", 15, 15); +return x_1; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__1; +x_2 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1; +x_3 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__2; +x_4 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +x_2 = lean_unsigned_to_nat(2u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(3); +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(3); +x_2 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4; +x_2 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_reduceDoNotSimp), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3; +x_3 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8; +x_4 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9; +x_5 = l_Lean_Meta_Simp_registerBuiltinDSimproc(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addDoNotSimp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3; +x_6 = 0; +x_7 = l_Lean_Meta_Simp_Simprocs_add(x_1, x_5, x_6, x_2, x_3, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addDoNotSimp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Meta_Grind_addDoNotSimp(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_1); +x_6 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_6, 0, x_5); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_6, 0, x_1); +x_7 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_5); +return x_8; +} +} +static lean_object* _init_l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1; +lean_inc(x_1); +x_6 = l_Lean_Expr_cleanupAnnotations(x_1); +x_7 = l_Lean_Expr_isApp(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_6); +x_8 = lean_box(0); +x_9 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(x_1, x_8, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_Expr_appArg(x_6, lean_box(0)); +x_11 = l_Lean_Expr_appFnCleanup(x_6, lean_box(0)); +x_12 = l_Lean_Expr_isApp(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); +x_13 = lean_box(0); +x_14 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(x_1, x_13, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); +x_16 = l_Lean_Meta_Grind_markAsDoNotSimp___closed__4; +x_17 = l_Lean_Expr_isConstOf(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_10); +x_18 = lean_box(0); +x_19 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(x_1, x_18, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; +} +else +{ +lean_object* x_20; +lean_dec(x_1); +x_20 = lean_apply_4(x_5, x_10, x_2, x_3, x_4); +return x_20; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_5, 0, x_1); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_eraseDoNotSimp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_eraseDoNotSimp___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Lean_Meta_Grind_eraseDoNotSimp___closed__1; +x_6 = l_Lean_Meta_Grind_eraseDoNotSimp___closed__2; +x_7 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_1, x_5, x_6, x_2, x_3, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Simproc(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Simp_Simproc(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_DoNotSimp(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Init_Grind_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Simproc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Simp_Simproc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_markAsDoNotSimp___closed__1 = _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsDoNotSimp___closed__1); +l_Lean_Meta_Grind_markAsDoNotSimp___closed__2 = _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsDoNotSimp___closed__2); +l_Lean_Meta_Grind_markAsDoNotSimp___closed__3 = _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsDoNotSimp___closed__3); +l_Lean_Meta_Grind_markAsDoNotSimp___closed__4 = _init_l_Lean_Meta_Grind_markAsDoNotSimp___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsDoNotSimp___closed__4); +l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_reduceDoNotSimp___lambda__2___closed__1); +l_Lean_Meta_Grind_reduceDoNotSimp___closed__1 = _init_l_Lean_Meta_Grind_reduceDoNotSimp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_reduceDoNotSimp___closed__1); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__1); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__2); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__3); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__4); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__5); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__6); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__7); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__8); +l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9 = _init_l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153____closed__9); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_reduceDoNotSimp_declare__1____x40_Lean_Meta_Tactic_Grind_DoNotSimp___hyg_153_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3___closed__1); +l_Lean_Meta_Grind_eraseDoNotSimp___closed__1 = _init_l_Lean_Meta_Grind_eraseDoNotSimp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_eraseDoNotSimp___closed__1); +l_Lean_Meta_Grind_eraseDoNotSimp___closed__2 = _init_l_Lean_Meta_Grind_eraseDoNotSimp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_eraseDoNotSimp___closed__2); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c index b573249175a5..98505b946360 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatch.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.EMatch -// Imports: Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Intro +// Imports: Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.DoNotSimp #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,7 +14,7 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -23,6 +23,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5; +lean_object* l_Lean_mkNatLit(lean_object*); static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); @@ -33,22 +35,31 @@ lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, l lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___boxed(lean_object**); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2; +lean_object* l_Lean_Meta_isOffset_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___boxed(lean_object**); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__6; uint64_t lean_uint64_lor(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; uint8_t l_Lean_Expr_isApp(lean_object*); +lean_object* l_Lean_Meta_Grind_markAsDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_trySynthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3; lean_object* l_Lean_MessageData_ofList(lean_object*); -lean_object* l_Lean_Meta_Grind_iterate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); @@ -60,67 +71,74 @@ static lean_object* l_Lean_Meta_Grind_ematch___closed__1; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Meta_Match_isMatchEqnTheorem(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4; lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1; lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1; extern lean_object* l_Lean_Meta_Grind_grind_debug_proofs; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorems___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3; lean_object* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_assertNext_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_evalNat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1; -static lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3; uint8_t l_Lean_Expr_isBVar(lean_object*); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; lean_object* lean_checked_assign(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_reverse___rarg(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__3; @@ -128,101 +146,119 @@ uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedState; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1(size_t, size_t, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8; +extern lean_object* l___private_Lean_Expr_0__Lean_natAddFn; lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2; lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ematch___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7; static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forMAux___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__3; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp(lean_object*); lean_object* l_Lean_Expr_toHeadIndex(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(lean_object*); +lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f(lean_object*); static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constName_x21(lean_object*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___boxed(lean_object**); lean_object* l_Lean_MessageData_ofConst(lean_object*); +static lean_object* l_Lean_Meta_Grind_ematchAndAssert___closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedCnstr; uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1; static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedChoice___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1___boxed(lean_object**); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2; uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); static uint64_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2; +lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions___boxed(lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3; lean_object* l_Lean_Expr_appFn_x21(lean_object*); +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ematch___closed__3; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4; extern lean_object* l_Lean_Meta_instMonadMetaM; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9___boxed(lean_object**); static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); lean_object* l_OptionT_instMonad___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___boxed(lean_object**); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_tail_x21___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__6; @@ -230,47 +266,56 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_e LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_ematchTheorems(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallMetaBoundedTelescope(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Meta_Grind_ENode_isCongrRoot(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3; lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; +lean_object* l_Lean_Expr_constLevels_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6___boxed(lean_object**); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1___boxed(lean_object**); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7___boxed(lean_object**); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1; +lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2; lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_isEqv(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__3; -static lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); @@ -278,48 +323,50 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_ static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1; LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_left(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___boxed(lean_object**); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Expr_fvar___override(lean_object*); size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__1; -static size_t l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7; lean_object* lean_string_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3; +lean_object* l_Lean_Meta_Grind_Origin_key(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -329,10 +376,12 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchThe lean_object* l_Lean_Expr_bvarIdx_x21(lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConst(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; uint8_t l_Lean_Expr_isFVar(lean_object*); @@ -342,18 +391,24 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_expr_instantiate1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_assertNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; lean_object* l_ReaderT_instMonad___rarg(lean_object*); +lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_ematchAndAssert___closed__2; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1() { _start: { @@ -586,13 +641,15 @@ return x_6; static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__3() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 0; x_2 = l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2; -x_3 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_1); -return x_3; +x_3 = l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; +x_4 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); +return x_4; } } static lean_object* _init_l_Lean_Meta_Grind_EMatch_instInhabitedContext() { @@ -614,13 +671,15 @@ return x_1; static lean_object* _init_l_Lean_Meta_Grind_EMatch_M_run_x27___rarg___closed__1() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 1; x_2 = l_Lean_Meta_Grind_EMatch_instInhabitedContext___closed__2; -x_3 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_1); -return x_3; +x_3 = l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1; +x_4 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); +return x_4; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatch_M_run_x27___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { @@ -701,6 +760,44 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_EMatch_M_run_x27___rarg), 10, return x_2; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_3); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_3, 1); +lean_dec(x_15); +lean_ctor_set(x_3, 1, x_1); +x_16 = lean_apply_11(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_16; +} +else +{ +uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); +x_18 = lean_ctor_get(x_3, 0); +lean_inc(x_18); +lean_dec(x_3); +x_19 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_1); +lean_ctor_set_uint8(x_19, sizeof(void*)*2, x_17); +x_20 = lean_apply_11(x_2, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_withInitApp___rarg), 13, 0); +return x_2; +} +} static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1___closed__1() { _start: { @@ -796,7 +893,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__2; -x_3 = lean_unsigned_to_nat(80u); +x_3 = lean_unsigned_to_nat(88u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1063,88 +1160,106 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1() { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = l_Lean_Expr_appArg_x21(x_1); -x_15 = l_Lean_Expr_appArg_x21(x_2); -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_17 = lean_expr_eqv(x_14, x_16); -if (x_17 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2() { +_start: { -uint8_t x_18; -x_18 = l_Lean_Expr_isBVar(x_14); -if (x_18 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("!isPatternDontCare pArg\n ", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3() { +_start: { -lean_object* x_19; -x_19 = l_Lean_Meta_Grind_groundPattern_x3f(x_14); -if (lean_obj_tag(x_19) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4() { +_start: { -lean_object* x_20; uint8_t x_21; -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_14); -lean_ctor_set(x_20, 1, x_15); -x_21 = !lean_is_exclusive(x_3); -if (x_21 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatch.0.Lean.Meta.Grind.EMatch.matchArg\?", 73, 73); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5() { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_3, 0); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 1, x_22); -lean_ctor_set(x_3, 0, x_23); -x_24 = l_Lean_Expr_appFn_x21(x_1); -x_25 = l_Lean_Expr_appFn_x21(x_2); -x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_24, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_25); -lean_dec(x_24); -return x_26; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4; +x_3 = lean_unsigned_to_nat(109u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6() { +_start: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_27 = lean_ctor_get(x_3, 0); -x_28 = lean_ctor_get(x_3, 1); -x_29 = lean_ctor_get(x_3, 2); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_3); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_20); -lean_ctor_set(x_30, 1, x_27); -x_31 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_28); -lean_ctor_set(x_31, 2, x_29); -x_32 = l_Lean_Expr_appFn_x21(x_1); -x_33 = l_Lean_Expr_appFn_x21(x_2); -x_34 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_31, x_32, x_33, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_33); -lean_dec(x_32); -return x_34; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.isNone <| isOffsetPattern\? pArg\n ", 43, 43); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7() { +_start: { -lean_object* x_35; lean_object* x_36; -lean_dec(x_14); -x_35 = lean_ctor_get(x_19, 0); -lean_inc(x_35); -lean_dec(x_19); -x_36 = l_Lean_Meta_Grind_isEqv(x_35, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_36) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8() { +_start: { -lean_object* x_37; uint8_t x_38; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_unbox(x_37); -lean_dec(x_37); -if (x_38 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4; +x_3 = lean_unsigned_to_nat(108u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_39; -lean_dec(x_12); +lean_object* x_13; uint8_t x_14; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_14 = lean_expr_eqv(x_2, x_13); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = l_Lean_Expr_isBVar(x_2); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = l_Lean_Meta_Grind_groundPattern_x3f(x_2); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +lean_inc(x_2); +x_17 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_2); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1152,48 +1267,75 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -x_39 = !lean_is_exclusive(x_36); -if (x_39 == 0) +lean_dec(x_4); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_2); +lean_ctor_set(x_18, 1, x_3); +x_19 = !lean_is_exclusive(x_1); +if (x_19 == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_36, 0); -lean_dec(x_40); -x_41 = lean_box(0); -lean_ctor_set(x_36, 0, x_41); -return x_36; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_1, 0); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_1, 0, x_21); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_1); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_12); +return x_23; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_36, 1); -lean_inc(x_42); -lean_dec(x_36); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -return x_44; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_ctor_get(x_1, 1); +x_26 = lean_ctor_get(x_1, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_1); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_18); +lean_ctor_set(x_27, 1, x_24); +x_28 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_25); +lean_ctor_set(x_28, 2, x_26); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_12); +return x_30; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_36, 1); -lean_inc(x_45); -lean_dec(x_36); -x_46 = l_Lean_Expr_appFn_x21(x_1); -x_47 = l_Lean_Expr_appFn_x21(x_2); -x_48 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_46, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_45); -lean_dec(x_47); -lean_dec(x_46); -return x_48; -} -} -else +uint8_t x_31; +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_17); +if (x_31 == 0) { -uint8_t x_49; -lean_dec(x_12); +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_17, 0); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_32, 0); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +x_36 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_34); +if (lean_obj_tag(x_36) == 0) +{ +uint8_t x_37; +x_37 = lean_expr_eqv(x_34, x_13); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1201,52 +1343,96 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -x_49 = !lean_is_exclusive(x_36); -if (x_49 == 0) +lean_dec(x_4); +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_35); +lean_ctor_set(x_38, 2, x_3); +x_39 = !lean_is_exclusive(x_1); +if (x_39 == 0) { -return x_36; +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_1, 0); +lean_ctor_set_tag(x_32, 1); +lean_ctor_set(x_32, 1, x_40); +lean_ctor_set(x_32, 0, x_38); +lean_ctor_set(x_1, 0, x_32); +lean_ctor_set(x_17, 0, x_1); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_17); +lean_ctor_set(x_41, 1, x_12); +return x_41; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_36, 0); -x_51 = lean_ctor_get(x_36, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_36); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_42 = lean_ctor_get(x_1, 0); +x_43 = lean_ctor_get(x_1, 1); +x_44 = lean_ctor_get(x_1, 2); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_1); +lean_ctor_set_tag(x_32, 1); +lean_ctor_set(x_32, 1, x_42); +lean_ctor_set(x_32, 0, x_38); +x_45 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_45, 0, x_32); +lean_ctor_set(x_45, 1, x_43); +lean_ctor_set(x_45, 2, x_44); +lean_ctor_set(x_17, 0, x_45); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_17); +lean_ctor_set(x_46, 1, x_12); +return x_46; } } +else +{ +lean_object* x_47; lean_object* x_48; +lean_free_object(x_32); +lean_dec(x_35); +lean_dec(x_34); +lean_free_object(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_47 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5; +x_48 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_47, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_48; } } else { -lean_object* x_53; lean_object* x_54; -x_53 = l_Lean_Expr_bvarIdx_x21(x_14); -lean_dec(x_14); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_54 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f(x_3, x_53, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_53); -if (lean_obj_tag(x_54) == 0) +lean_object* x_49; lean_object* x_50; +lean_dec(x_36); +lean_free_object(x_32); +lean_dec(x_35); +lean_dec(x_34); +lean_free_object(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8; +x_50 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_50; +} +} +else { -lean_object* x_55; -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -if (lean_obj_tag(x_55) == 0) +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_32, 0); +x_52 = lean_ctor_get(x_32, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_32); +lean_inc(x_51); +x_53 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_51); +if (lean_obj_tag(x_53) == 0) { -uint8_t x_56; -lean_dec(x_12); +uint8_t x_54; +x_54 = lean_expr_eqv(x_51, x_13); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1254,50 +1440,98 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_56 = !lean_is_exclusive(x_54); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_54, 0); -lean_dec(x_57); -x_58 = lean_box(0); -lean_ctor_set(x_54, 0, x_58); -return x_54; +lean_dec(x_4); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_51); +lean_ctor_set(x_55, 1, x_52); +lean_ctor_set(x_55, 2, x_3); +x_56 = lean_ctor_get(x_1, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_1, 1); +lean_inc(x_57); +x_58 = lean_ctor_get(x_1, 2); +lean_inc(x_58); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + x_59 = x_1; +} else { + lean_dec_ref(x_1); + x_59 = lean_box(0); } -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_54, 1); -lean_inc(x_59); -lean_dec(x_54); -x_60 = lean_box(0); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_59); -return x_61; +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_55); +lean_ctor_set(x_60, 1, x_56); +if (lean_is_scalar(x_59)) { + x_61 = lean_alloc_ctor(0, 3, 0); +} else { + x_61 = x_59; } +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_57); +lean_ctor_set(x_61, 2, x_58); +lean_ctor_set(x_17, 0, x_61); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_17); +lean_ctor_set(x_62, 1, x_12); +return x_62; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_62 = lean_ctor_get(x_54, 1); -lean_inc(x_62); -lean_dec(x_54); -x_63 = lean_ctor_get(x_55, 0); -lean_inc(x_63); -lean_dec(x_55); -x_64 = l_Lean_Expr_appFn_x21(x_1); -x_65 = l_Lean_Expr_appFn_x21(x_2); -x_66 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_63, x_64, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_62); -lean_dec(x_65); -lean_dec(x_64); +lean_object* x_63; lean_object* x_64; +lean_dec(x_52); +lean_dec(x_51); +lean_free_object(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5; +x_64 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_64; +} +} +else +{ +lean_object* x_65; lean_object* x_66; +lean_dec(x_53); +lean_dec(x_52); +lean_dec(x_51); +lean_free_object(x_17); +lean_dec(x_3); +lean_dec(x_1); +x_65 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8; +x_66 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_65, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_66; } } +} else { -uint8_t x_67; -lean_dec(x_12); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_67 = lean_ctor_get(x_17, 0); +lean_inc(x_67); +lean_dec(x_17); +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_70 = x_67; +} else { + lean_dec_ref(x_67); + x_70 = lean_box(0); +} +lean_inc(x_68); +x_71 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_68); +if (lean_obj_tag(x_71) == 0) +{ +uint8_t x_72; +x_72 = lean_expr_eqv(x_68, x_13); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1305,49 +1539,88 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_67 = !lean_is_exclusive(x_54); -if (x_67 == 0) -{ -return x_54; +lean_dec(x_4); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_68); +lean_ctor_set(x_73, 1, x_69); +lean_ctor_set(x_73, 2, x_3); +x_74 = lean_ctor_get(x_1, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_1, 1); +lean_inc(x_75); +x_76 = lean_ctor_get(x_1, 2); +lean_inc(x_76); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + x_77 = x_1; +} else { + lean_dec_ref(x_1); + x_77 = lean_box(0); } -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_54, 0); -x_69 = lean_ctor_get(x_54, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_54); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +if (lean_is_scalar(x_70)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_70; + lean_ctor_set_tag(x_78, 1); +} +lean_ctor_set(x_78, 0, x_73); +lean_ctor_set(x_78, 1, x_74); +if (lean_is_scalar(x_77)) { + x_79 = lean_alloc_ctor(0, 3, 0); +} else { + x_79 = x_77; } +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_75); +lean_ctor_set(x_79, 2, x_76); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_12); +return x_81; } +else +{ +lean_object* x_82; lean_object* x_83; +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_3); +lean_dec(x_1); +x_82 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5; +x_83 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_82, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_83; } } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -lean_dec(x_15); -lean_dec(x_14); -x_71 = l_Lean_Expr_appFn_x21(x_1); -x_72 = l_Lean_Expr_appFn_x21(x_2); -x_73 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_3, x_71, x_72, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_72); +lean_object* x_84; lean_object* x_85; lean_dec(x_71); -return x_73; +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_3); +lean_dec(x_1); +x_84 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8; +x_85 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___spec__1(x_84, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_85; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +} +else { -uint8_t x_13; -x_13 = l_Lean_Expr_isApp(x_2); -if (x_13 == 0) +uint8_t x_86; +lean_dec(x_2); +x_86 = !lean_is_exclusive(x_16); +if (x_86 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_16, 0); +x_88 = l_Lean_Meta_Grind_isEqv(x_87, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1356,694 +1629,604 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_12); -return x_15; -} -else +if (lean_obj_tag(x_88) == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_box(0); -x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_2, x_3, x_1, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_17; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_object* x_89; uint8_t x_90; +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_unbox(x_89); +lean_dec(x_89); +if (x_90 == 0) { -lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_4); -lean_dec(x_2); +uint8_t x_91; +lean_free_object(x_16); lean_dec(x_1); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_3); -lean_dec(x_2); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; -x_15 = l_Lean_Meta_Grind_getNext(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_1); -if (x_18 == 0) +x_91 = !lean_is_exclusive(x_88); +if (x_91 == 0) { -lean_object* x_19; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_15, 0, x_19); -return x_15; +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_88, 0); +lean_dec(x_92); +x_93 = lean_box(0); +lean_ctor_set(x_88, 0, x_93); +return x_88; } else { -lean_object* x_20; -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_17); -lean_ctor_set(x_15, 0, x_20); -return x_15; +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_88, 1); +lean_inc(x_94); +lean_dec(x_88); +x_95 = lean_box(0); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_94); +return x_96; } } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_1); -if (x_23 == 0) +uint8_t x_97; +x_97 = !lean_is_exclusive(x_88); +if (x_97 == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_22); -return x_25; +lean_object* x_98; +x_98 = lean_ctor_get(x_88, 0); +lean_dec(x_98); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_88, 0, x_16); +return x_88; } else { -lean_object* x_26; lean_object* x_27; -x_26 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_26, 0, x_21); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_22); -return x_27; +lean_object* x_99; lean_object* x_100; +x_99 = lean_ctor_get(x_88, 1); +lean_inc(x_99); +lean_dec(x_88); +lean_ctor_set(x_16, 0, x_1); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_16); +lean_ctor_set(x_100, 1, x_99); +return x_100; } } } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_15); -if (x_28 == 0) +uint8_t x_101; +lean_free_object(x_16); +lean_dec(x_1); +x_101 = !lean_is_exclusive(x_88); +if (x_101 == 0) { -return x_15; +return x_88; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_15, 0); -x_30 = lean_ctor_get(x_15, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_15); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; -} +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_88, 0); +x_103 = lean_ctor_get(x_88, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_88); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; } } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { -_start: +else { -lean_object* x_19; lean_object* x_35; -lean_inc(x_7); -x_35 = l_Lean_Meta_Grind_getENode(x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_35) == 0) +lean_object* x_105; lean_object* x_106; +x_105 = lean_ctor_get(x_16, 0); +lean_inc(x_105); +lean_dec(x_16); +x_106 = l_Lean_Meta_Grind_isEqv(x_105, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +if (lean_obj_tag(x_106) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_109; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = lean_ctor_get(x_36, 8); -lean_inc(x_38); -x_109 = lean_nat_dec_le(x_38, x_4); -if (x_109 == 0) +lean_object* x_107; uint8_t x_108; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_unbox(x_107); +lean_dec(x_107); +if (x_108 == 0) { -lean_object* x_110; lean_object* x_111; -lean_dec(x_38); -lean_dec(x_36); -x_110 = lean_box(0); -x_111 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_110, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_111; -goto block_34; +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_dec(x_1); +x_109 = lean_ctor_get(x_106, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_110 = x_106; +} else { + lean_dec_ref(x_106); + x_110 = lean_box(0); +} +x_111 = lean_box(0); +if (lean_is_scalar(x_110)) { + x_112 = lean_alloc_ctor(0, 2, 0); +} else { + x_112 = x_110; +} +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_109); +return x_112; } else { -uint8_t x_112; -x_112 = lean_ctor_get_uint8(x_36, sizeof(void*)*10 + 4); -if (x_112 == 0) -{ -lean_object* x_113; uint8_t x_114; -x_113 = lean_ctor_get(x_36, 3); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_113 = lean_ctor_get(x_106, 1); lean_inc(x_113); -lean_dec(x_36); -x_114 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_7, x_113); -lean_dec(x_113); -if (x_114 == 0) -{ -lean_object* x_115; lean_object* x_116; -lean_dec(x_38); -x_115 = lean_box(0); -x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_115, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_116; -goto block_34; +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_114 = x_106; +} else { + lean_dec_ref(x_106); + x_114 = lean_box(0); +} +x_115 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_115, 0, x_1); +if (lean_is_scalar(x_114)) { + x_116 = lean_alloc_ctor(0, 2, 0); +} else { + x_116 = x_114; +} +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_113); +return x_116; } -else -{ -lean_object* x_117; uint8_t x_118; -x_117 = l_Lean_Expr_getAppFn(x_7); -x_118 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_5, x_117); -lean_dec(x_117); -if (x_118 == 0) -{ -lean_object* x_119; lean_object* x_120; -lean_dec(x_38); -x_119 = lean_box(0); -x_120 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_119, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_120; -goto block_34; } else { -lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_121 = lean_unsigned_to_nat(0u); -x_122 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_7, x_121); -x_123 = lean_nat_dec_eq(x_122, x_6); -lean_dec(x_122); -if (x_123 == 0) -{ -lean_object* x_124; lean_object* x_125; -lean_dec(x_38); -x_124 = lean_box(0); -x_125 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_124, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_125; -goto block_34; +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_dec(x_1); +x_117 = lean_ctor_get(x_106, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_106, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_119 = x_106; +} else { + lean_dec_ref(x_106); + x_119 = lean_box(0); } -else -{ -lean_object* x_126; -x_126 = lean_box(0); -x_39 = x_126; -goto block_108; +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 2, 0); +} else { + x_120 = x_119; } +lean_ctor_set(x_120, 0, x_117); +lean_ctor_set(x_120, 1, x_118); +return x_120; } } } -else -{ -lean_object* x_127; uint8_t x_128; -lean_dec(x_36); -x_127 = l_Lean_Expr_getAppFn(x_7); -x_128 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_5, x_127); -lean_dec(x_127); -if (x_128 == 0) -{ -lean_object* x_129; lean_object* x_130; -lean_dec(x_38); -x_129 = lean_box(0); -x_130 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_129, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_130; -goto block_34; } else { -lean_object* x_131; lean_object* x_132; uint8_t x_133; -x_131 = lean_unsigned_to_nat(0u); -x_132 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_7, x_131); -x_133 = lean_nat_dec_eq(x_132, x_6); -lean_dec(x_132); -if (x_133 == 0) -{ -lean_object* x_134; lean_object* x_135; -lean_dec(x_38); -x_134 = lean_box(0); -x_135 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_134, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -x_19 = x_135; -goto block_34; +lean_object* x_121; lean_object* x_122; +x_121 = l_Lean_Expr_bvarIdx_x21(x_2); +lean_dec(x_2); +x_122 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f(x_1, x_121, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_121); +return x_122; +} } else { -lean_object* x_136; -x_136 = lean_box(0); -x_39 = x_136; -goto block_108; -} -} +lean_object* x_123; lean_object* x_124; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_123 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_123, 0, x_1); +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_12); +return x_124; } } -block_108: -{ -lean_object* x_40; -lean_dec(x_39); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_1); -x_40 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_37); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_dec(x_38); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_box(0); -x_44 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_43, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_42); -x_19 = x_44; -goto block_34; } -else -{ -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = lean_ctor_get(x_41, 0); -lean_inc(x_45); -lean_dec(x_41); -x_46 = lean_ctor_get(x_40, 1); -lean_inc(x_46); -lean_dec(x_40); -x_47 = !lean_is_exclusive(x_45); -if (x_47 == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_48; uint8_t x_49; -x_48 = lean_ctor_get(x_45, 1); -x_49 = lean_nat_dec_le(x_38, x_48); -if (x_49 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) { -lean_object* x_50; uint8_t x_51; -lean_dec(x_48); -lean_ctor_set(x_45, 1, x_38); -x_50 = lean_st_ref_take(x_9, x_46); -x_51 = !lean_is_exclusive(x_50); -if (x_51 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_nat_dec_le(x_2, x_4); +if (x_5 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_50, 0); -x_53 = lean_ctor_get(x_50, 1); -lean_ctor_set_tag(x_50, 1); -lean_ctor_set(x_50, 1, x_52); -lean_ctor_set(x_50, 0, x_45); -x_54 = lean_st_ref_set(x_9, x_50, x_53); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = lean_box(0); -x_57 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_55); -x_19 = x_57; -goto block_34; +lean_dec(x_4); +lean_ctor_set(x_1, 1, x_2); +return x_1; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_58 = lean_ctor_get(x_50, 0); -x_59 = lean_ctor_get(x_50, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_50); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_45); -lean_ctor_set(x_60, 1, x_58); -x_61 = lean_st_ref_set(x_9, x_60, x_59); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -lean_dec(x_61); -x_63 = lean_box(0); -x_64 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_63, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_62); -x_19 = x_64; -goto block_34; +lean_dec(x_2); +return x_1; } } else { -lean_object* x_65; uint8_t x_66; -lean_dec(x_38); -x_65 = lean_st_ref_take(x_9, x_46); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_1); +x_9 = lean_nat_dec_le(x_2, x_7); +if (x_9 == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_67 = lean_ctor_get(x_65, 0); -x_68 = lean_ctor_get(x_65, 1); -lean_ctor_set_tag(x_65, 1); -lean_ctor_set(x_65, 1, x_67); -lean_ctor_set(x_65, 0, x_45); -x_69 = lean_st_ref_set(x_9, x_65, x_68); -x_70 = lean_ctor_get(x_69, 1); -lean_inc(x_70); -lean_dec(x_69); -x_71 = lean_box(0); -x_72 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_71, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_70); -x_19 = x_72; -goto block_34; +lean_object* x_10; +lean_dec(x_7); +x_10 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 2, x_8); +return x_10; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_73 = lean_ctor_get(x_65, 0); -x_74 = lean_ctor_get(x_65, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_65); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_45); -lean_ctor_set(x_75, 1, x_73); -x_76 = lean_st_ref_set(x_9, x_75, x_74); -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -lean_dec(x_76); -x_78 = lean_box(0); -x_79 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_78, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_77); -x_19 = x_79; -goto block_34; +lean_object* x_11; +lean_dec(x_2); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_6); +lean_ctor_set(x_11, 1, x_7); +lean_ctor_set(x_11, 2, x_8); +return x_11; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; -x_80 = lean_ctor_get(x_45, 0); -x_81 = lean_ctor_get(x_45, 1); -x_82 = lean_ctor_get(x_45, 2); -lean_inc(x_82); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_45); -x_83 = lean_nat_dec_le(x_38, x_81); -if (x_83 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_take(x_3, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_81); -x_84 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_84, 0, x_80); -lean_ctor_set(x_84, 1, x_38); -lean_ctor_set(x_84, 2, x_82); -x_85 = lean_st_ref_take(x_9, x_46); -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_85, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_88 = x_85; -} else { - lean_dec_ref(x_85); - x_88 = lean_box(0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 1, x_15); +lean_ctor_set(x_13, 0, x_1); +x_17 = lean_st_ref_set(x_3, x_13, x_16); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_17, 0, x_20); +return x_17; } -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); -} else { - x_89 = x_88; - lean_ctor_set_tag(x_89, 1); +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; } -lean_ctor_set(x_89, 0, x_84); -lean_ctor_set(x_89, 1, x_86); -x_90 = lean_st_ref_set(x_9, x_89, x_87); -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec(x_90); -x_92 = lean_box(0); -x_93 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_92, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_91); -x_19 = x_93; -goto block_34; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -lean_dec(x_38); -x_94 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_94, 0, x_80); -lean_ctor_set(x_94, 1, x_81); -lean_ctor_set(x_94, 2, x_82); -x_95 = lean_st_ref_take(x_9, x_46); -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_98 = x_95; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_24); +x_27 = lean_st_ref_set(x_3, x_26, x_25); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_29 = x_27; } else { - lean_dec_ref(x_95); - x_98 = lean_box(0); + lean_dec_ref(x_27); + x_29 = lean_box(0); } -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(1, 2, 0); +x_30 = lean_box(0); +if (lean_is_scalar(x_29)) { + x_31 = lean_alloc_ctor(0, 2, 0); } else { - x_99 = x_98; - lean_ctor_set_tag(x_99, 1); + x_31 = x_29; } -lean_ctor_set(x_99, 0, x_94); -lean_ctor_set(x_99, 1, x_96); -x_100 = lean_st_ref_set(x_9, x_99, x_97); -x_101 = lean_ctor_get(x_100, 1); -lean_inc(x_101); -lean_dec(x_100); -x_102 = lean_box(0); -x_103 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_102, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_101); -x_19 = x_103; -goto block_34; +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +return x_31; } } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_104; -lean_dec(x_38); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = l_Lean_Expr_appArg_x21(x_1); +x_15 = l_Lean_Expr_appArg_x21(x_2); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_3, x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_104 = !lean_is_exclusive(x_40); -if (x_104 == 0) +lean_dec(x_6); +lean_dec(x_5); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) { -x_19 = x_40; -goto block_34; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_16, 0, x_20); +return x_16; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_105 = lean_ctor_get(x_40, 0); -x_106 = lean_ctor_get(x_40, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_40); -x_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -x_19 = x_107; -goto block_34; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_dec(x_16); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; } } +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_dec(x_16); +x_25 = lean_ctor_get(x_17, 0); +lean_inc(x_25); +lean_dec(x_17); +x_26 = l_Lean_Expr_appFn_x21(x_1); +x_27 = l_Lean_Expr_appFn_x21(x_2); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_25, x_26, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_24); +lean_dec(x_27); +lean_dec(x_26); +return x_28; } } else { -uint8_t x_137; +uint8_t x_29; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_137 = !lean_is_exclusive(x_35); -if (x_137 == 0) +lean_dec(x_6); +lean_dec(x_5); +x_29 = !lean_is_exclusive(x_16); +if (x_29 == 0) { -x_19 = x_35; -goto block_34; +return x_16; } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_138 = lean_ctor_get(x_35, 0); -x_139 = lean_ctor_get(x_35, 1); -lean_inc(x_139); -lean_inc(x_138); -lean_dec(x_35); -x_140 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_140, 0, x_138); -lean_ctor_set(x_140, 1, x_139); -x_19 = x_140; -goto block_34; +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_16, 0); +x_31 = lean_ctor_get(x_16, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_16); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } -block_34: -{ -if (lean_obj_tag(x_19) == 0) +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -if (lean_obj_tag(x_20) == 0) +uint8_t x_13; +x_13 = l_Lean_Expr_isApp(x_2); +if (x_13 == 0) { -uint8_t x_21; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_14; lean_object* x_15; lean_dec(x_11); lean_dec(x_10); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_19, 0); -lean_dec(x_22); -x_23 = lean_ctor_get(x_20, 0); -lean_inc(x_23); -lean_dec(x_20); -lean_ctor_set(x_19, 0, x_23); -return x_19; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_1); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_12); +return x_15; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_ctor_get(x_20, 0); -lean_inc(x_25); -lean_dec(x_20); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; -} +lean_object* x_16; lean_object* x_17; +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_2, x_3, x_1, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_17; } -else -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_19, 1); -lean_inc(x_27); -lean_dec(x_19); -x_28 = lean_ctor_get(x_20, 0); -lean_inc(x_28); -lean_dec(x_20); -x_7 = x_28; -x_18 = x_27; -goto _start; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_30; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +lean_dec(x_2); lean_dec(x_1); -x_30 = !lean_is_exclusive(x_19); -if (x_30 == 0) -{ -return x_19; +return x_14; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_19, 0); -x_32 = lean_ctor_get(x_19, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_19); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +lean_dec(x_2); +return x_13; } } +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Lean_Meta_Grind_getNext(x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_1); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_15, 0, x_19); +return x_15; } +else +{ +lean_object* x_20; +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_15, 0, x_20); +return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_15 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_8, x_9, x_10, x_11, x_12, x_13, x_14); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); lean_dec(x_15); -x_18 = l_Lean_Expr_getAppFn(x_2); -x_19 = lean_unsigned_to_nat(0u); -x_20 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_19); -lean_inc(x_3); -x_21 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_16, x_18, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -lean_dec(x_20); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_3); -if (lean_obj_tag(x_21) == 0) -{ -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_1); +if (x_23 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -x_24 = lean_box(0); -lean_ctor_set(x_21, 0, x_24); -return x_21; +lean_object* x_24; lean_object* x_25; +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_22); +return x_25; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = lean_box(0); +lean_object* x_26; lean_object* x_27; +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_21); x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); +lean_ctor_set(x_27, 1, x_22); return x_27; } } +} else { uint8_t x_28; -x_28 = !lean_is_exclusive(x_21); +x_28 = !lean_is_exclusive(x_15); if (x_28 == 0) { -return x_21; +return x_15; } else { lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_21, 0); -x_30 = lean_ctor_get(x_21, 1); +x_29 = lean_ctor_get(x_15, 0); +x_30 = lean_ctor_get(x_15, 1); lean_inc(x_30); lean_inc(x_29); -lean_dec(x_21); +lean_dec(x_15); x_31 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); @@ -2052,563 +2235,577 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_object* x_15; -x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_19; lean_object* x_20; lean_object* x_26; +lean_inc(x_7); +x_26 = l_Lean_Meta_Grind_getENode(x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_81; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_ctor_get(x_27, 8); +lean_inc(x_29); +x_81 = lean_nat_dec_lt(x_29, x_4); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; +lean_dec(x_29); +lean_dec(x_27); +x_82 = lean_box(0); +x_83 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_82, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_83) == 0) +{ +lean_object* x_84; lean_object* x_85; +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_19 = x_84; +x_20 = x_85; +goto block_25; +} +else +{ +uint8_t x_86; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_1); -return x_15; +x_86 = !lean_is_exclusive(x_83); +if (x_86 == 0) +{ +return x_83; } +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_83, 0); +x_88 = lean_ctor_get(x_83, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_83); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -_start: +} +} +else { -lean_object* x_19; -x_19 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_19; +uint8_t x_90; +x_90 = lean_ctor_get_uint8(x_27, sizeof(void*)*10 + 4); +if (x_90 == 0) +{ +uint8_t x_91; +x_91 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_27); +lean_dec(x_27); +if (x_91 == 0) +{ +lean_object* x_92; lean_object* x_93; +lean_dec(x_29); +x_92 = lean_box(0); +x_93 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_92, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_93) == 0) +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +x_19 = x_94; +x_20 = x_95; +goto block_25; } +else +{ +uint8_t x_96; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_96 = !lean_is_exclusive(x_93); +if (x_96 == 0) +{ +return x_93; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +else { -lean_object* x_15; -x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -return x_15; +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_93, 0); +x_98 = lean_ctor_get(x_93, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_93); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_100; +x_100 = lean_box(0); +x_30 = x_100; +goto block_80; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { -_start: +else { -if (lean_obj_tag(x_7) == 0) +lean_object* x_101; +lean_dec(x_27); +x_101 = lean_box(0); +x_30 = x_101; +goto block_80; +} +} +block_80: { -lean_object* x_21; -lean_dec(x_19); -lean_dec(x_18); +lean_object* x_31; uint8_t x_32; +lean_dec(x_30); +x_31 = l_Lean_Expr_getAppFn(x_7); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_eqvFunctions(x_5, x_31); +lean_dec(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_29); +x_33 = lean_box(0); +x_34 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_33, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_19 = x_35; +x_20 = x_36; +goto block_25; +} +else +{ +uint8_t x_37; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_8); -lean_ctor_set(x_21, 1, x_20); -return x_21; +x_37 = !lean_is_exclusive(x_34); +if (x_37 == 0) +{ +return x_34; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_30; -lean_dec(x_8); -x_22 = lean_ctor_get(x_7, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_7, 1); -lean_inc(x_23); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - x_24 = x_7; -} else { - lean_dec_ref(x_7); - x_24 = lean_box(0); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_34, 0); +x_39 = lean_ctor_get(x_34, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_34); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} } -lean_inc(x_22); -x_30 = l_Lean_Meta_Grind_getENode(x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_83; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = lean_ctor_get(x_31, 8); -lean_inc(x_33); -x_83 = lean_nat_dec_le(x_33, x_4); -if (x_83 == 0) -{ -lean_object* x_84; -lean_dec(x_33); -lean_dec(x_31); -lean_dec(x_24); -lean_dec(x_22); -x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_84; -x_26 = x_32; -goto block_29; } else { -uint8_t x_85; -x_85 = lean_ctor_get_uint8(x_31, sizeof(void*)*10 + 4); -if (x_85 == 0) +lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_41 = lean_unsigned_to_nat(0u); +x_42 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_7, x_41); +x_43 = lean_nat_dec_eq(x_42, x_6); +lean_dec(x_42); +if (x_43 == 0) { -lean_object* x_86; uint8_t x_87; -x_86 = lean_ctor_get(x_31, 3); -lean_inc(x_86); -lean_dec(x_31); -x_87 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_86, x_22); -lean_dec(x_86); -if (x_87 == 0) +lean_object* x_44; lean_object* x_45; +lean_dec(x_29); +x_44 = lean_box(0); +x_45 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_44, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_88; -lean_dec(x_33); -lean_dec(x_24); -lean_dec(x_22); -x_88 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_88; -x_26 = x_32; -goto block_29; +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_19 = x_46; +x_20 = x_47; +goto block_25; } else { -lean_object* x_89; -x_89 = lean_box(0); -x_34 = x_89; -goto block_82; -} +uint8_t x_48; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) +{ +return x_45; } else { -lean_object* x_90; -lean_dec(x_31); -x_90 = lean_box(0); -x_34 = x_90; -goto block_82; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 0); +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_45); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } -block_82: +} +else { -lean_object* x_35; -lean_dec(x_34); -lean_inc(x_19); -lean_inc(x_18); +lean_object* x_52; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_1); -x_35 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_32); -lean_dec(x_22); -if (lean_obj_tag(x_35) == 0) +x_52 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_7, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -if (lean_obj_tag(x_36) == 0) +lean_object* x_53; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_37; lean_object* x_38; -lean_dec(x_33); -lean_dec(x_24); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_38; -x_26 = x_37; -goto block_29; +lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_29); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_box(0); +x_56 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_55, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_54); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_19 = x_57; +x_20 = x_58; +goto block_25; } else { -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = lean_ctor_get(x_36, 0); -lean_inc(x_39); -lean_dec(x_36); -x_40 = lean_ctor_get(x_35, 1); -lean_inc(x_40); -lean_dec(x_35); -x_41 = !lean_is_exclusive(x_39); -if (x_41 == 0) -{ -lean_object* x_42; uint8_t x_43; -x_42 = lean_ctor_get(x_39, 1); -x_43 = lean_nat_dec_le(x_33, x_42); -if (x_43 == 0) +uint8_t x_59; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_56); +if (x_59 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_42); -lean_ctor_set(x_39, 1, x_33); -x_44 = lean_st_ref_take(x_11, x_40); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -if (lean_is_scalar(x_24)) { - x_47 = lean_alloc_ctor(1, 2, 0); -} else { - x_47 = x_24; -} -lean_ctor_set(x_47, 0, x_39); -lean_ctor_set(x_47, 1, x_45); -x_48 = lean_st_ref_set(x_11, x_47, x_46); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec(x_48); -x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_50; -x_26 = x_49; -goto block_29; +return x_56; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_dec(x_33); -x_51 = lean_st_ref_take(x_11, x_40); -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -if (lean_is_scalar(x_24)) { - x_54 = lean_alloc_ctor(1, 2, 0); -} else { - x_54 = x_24; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_56, 0); +x_61 = lean_ctor_get(x_56, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_56); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; } -lean_ctor_set(x_54, 0, x_39); -lean_ctor_set(x_54, 1, x_52); -x_55 = lean_st_ref_set(x_11, x_54, x_53); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_57; -x_26 = x_56; -goto block_29; } } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_58 = lean_ctor_get(x_39, 0); -x_59 = lean_ctor_get(x_39, 1); -x_60 = lean_ctor_get(x_39, 2); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_39); -x_61 = lean_nat_dec_le(x_33, x_59); -if (x_61 == 0) -{ -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -lean_dec(x_59); -x_62 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_62, 0, x_58); -lean_ctor_set(x_62, 1, x_33); -lean_ctor_set(x_62, 2, x_60); -x_63 = lean_st_ref_take(x_11, x_40); -x_64 = lean_ctor_get(x_63, 0); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_63 = lean_ctor_get(x_52, 1); +lean_inc(x_63); +lean_dec(x_52); +x_64 = lean_ctor_get(x_53, 0); lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -if (lean_is_scalar(x_24)) { - x_66 = lean_alloc_ctor(1, 2, 0); -} else { - x_66 = x_24; -} -lean_ctor_set(x_66, 0, x_62); -lean_ctor_set(x_66, 1, x_64); -x_67 = lean_st_ref_set(x_11, x_66, x_65); -x_68 = lean_ctor_get(x_67, 1); +lean_dec(x_53); +x_65 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_64, x_29); +x_66 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_63); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); lean_inc(x_68); +lean_dec(x_66); +x_69 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_3, x_7, x_67, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_68); lean_dec(x_67); -x_69 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_69; -x_26 = x_68; -goto block_29; +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_19 = x_70; +x_20 = x_71; +goto block_25; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_33); -x_70 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_70, 0, x_58); -lean_ctor_set(x_70, 1, x_59); -lean_ctor_set(x_70, 2, x_60); -x_71 = lean_st_ref_take(x_11, x_40); -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); -lean_inc(x_73); -lean_dec(x_71); -if (lean_is_scalar(x_24)) { - x_74 = lean_alloc_ctor(1, 2, 0); -} else { - x_74 = x_24; +uint8_t x_72; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_72 = !lean_is_exclusive(x_69); +if (x_72 == 0) +{ +return x_69; } -lean_ctor_set(x_74, 0, x_70); -lean_ctor_set(x_74, 1, x_72); -x_75 = lean_st_ref_set(x_11, x_74, x_73); -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -lean_dec(x_75); -x_77 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; -x_25 = x_77; -x_26 = x_76; -goto block_29; +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_69, 0); +x_74 = lean_ctor_get(x_69, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_69); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } } } else { -uint8_t x_78; -lean_dec(x_33); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_19); -lean_dec(x_18); +uint8_t x_76; +lean_dec(x_29); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); lean_dec(x_1); -x_78 = !lean_is_exclusive(x_35); -if (x_78 == 0) +x_76 = !lean_is_exclusive(x_52); +if (x_76 == 0) { -return x_35; +return x_52; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_35, 0); -x_80 = lean_ctor_get(x_35, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_35); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -return x_81; +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_52, 0); +x_78 = lean_ctor_get(x_52, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_52); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} } } } } else { -uint8_t x_91; -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_19); -lean_dec(x_18); +uint8_t x_102; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); lean_dec(x_1); -x_91 = !lean_is_exclusive(x_30); -if (x_91 == 0) +x_102 = !lean_is_exclusive(x_26); +if (x_102 == 0) { -return x_30; +return x_26; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_30, 0); -x_93 = lean_ctor_get(x_30, 1); -lean_inc(x_93); -lean_inc(x_92); -lean_dec(x_30); -x_94 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_94, 0, x_92); -lean_ctor_set(x_94, 1, x_93); -return x_94; +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_26, 0); +x_104 = lean_ctor_get(x_26, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_26); +x_105 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 1, x_104); +return x_105; } } -block_29: +block_25: { -lean_object* x_27; -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_1); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +else +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +lean_dec(x_19); x_7 = x_23; -x_8 = x_27; -x_9 = lean_box(0); -x_20 = x_26; +x_18 = x_20; goto _start; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; uint8_t x_15; -x_14 = lean_st_ref_get(x_5, x_13); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_14, 1); -x_18 = lean_ctor_get(x_16, 4); -lean_inc(x_18); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_15 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Expr_getAppFn(x_2); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_19); +lean_inc(x_3); +x_21 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_16, x_18, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +lean_dec(x_20); +lean_dec(x_18); lean_dec(x_16); -lean_inc(x_2); -x_19 = l_Lean_Expr_toHeadIndex(x_2); -x_20 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_18, x_19); -lean_dec(x_19); -if (lean_obj_tag(x_20) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_21; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_box(0); -lean_ctor_set(x_14, 0, x_21); -return x_14; -} -else +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_free_object(x_14); -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_17); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); lean_dec(x_23); -x_26 = lean_box(0); -x_27 = lean_box(0); -lean_inc(x_22); -x_28 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_22, x_24, x_26, x_22, x_22, x_27, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_25); -lean_dec(x_24); -lean_dec(x_22); -lean_dec(x_2); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; -x_30 = lean_ctor_get(x_28, 0); -lean_dec(x_30); -lean_ctor_set(x_28, 0, x_27); -return x_28; +x_24 = lean_box(0); +lean_ctor_set(x_21, 0, x_24); +return x_21; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_27); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; } } else { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_28); -if (x_33 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_21); +if (x_28 == 0) { -return x_28; +return x_21; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_28, 0); -x_35 = lean_ctor_get(x_28, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_28); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_21, 0); +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_21); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } } -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = lean_ctor_get(x_14, 0); -x_38 = lean_ctor_get(x_14, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_14); -x_39 = lean_ctor_get(x_37, 4); -lean_inc(x_39); -lean_dec(x_37); -lean_inc(x_2); -x_40 = l_Lean_Expr_toHeadIndex(x_2); -x_41 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_39, x_40); -lean_dec(x_40); -if (lean_obj_tag(x_41) == 0) +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_42; lean_object* x_43; +lean_object* x_15; +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -2617,84 +2814,13 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_42 = lean_box(0); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_38); -return x_43; -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_44 = lean_ctor_get(x_41, 0); -lean_inc(x_44); -lean_dec(x_41); -x_45 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_38); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_box(0); -x_49 = lean_box(0); -lean_inc(x_44); -x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_44, x_46, x_48, x_44, x_44, x_49, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); -lean_dec(x_46); -lean_dec(x_44); -lean_dec(x_2); -if (lean_obj_tag(x_50) == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -if (lean_is_scalar(x_52)) { - x_53 = lean_alloc_ctor(0, 2, 0); -} else { - x_53 = x_52; -} -lean_ctor_set(x_53, 0, x_49); -lean_ctor_set(x_53, 1, x_51); -return x_53; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = lean_ctor_get(x_50, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_50, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_56 = x_50; -} else { - lean_dec_ref(x_50); - x_56 = lean_box(0); -} -if (lean_is_scalar(x_56)) { - x_57 = lean_alloc_ctor(1, 2, 0); -} else { - x_57 = x_56; -} -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_55); -return x_57; -} -} -} +return x_15; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -2713,46 +2839,2437 @@ lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; _start: { -lean_object* x_21; -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_19; +x_19 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_21; +return x_19; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_object* x_15; +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_14; +lean_dec(x_2); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -uint8_t x_13; -x_13 = l_Lean_Expr_hasMVar(x_1); -if (x_13 == 0) -{ -lean_object* x_14; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -else +lean_object* x_18; lean_object* x_34; +lean_inc(x_6); +x_34 = l_Lean_Meta_Grind_getENode(x_6, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_35, 8); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_nat_dec_lt(x_37, x_5); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_37); +x_39 = lean_box(0); +x_40 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_39, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_36); +x_18 = x_40; +goto block_33; +} +else +{ +lean_object* x_41; +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_6); +x_41 = l_Lean_Meta_isOffset_x3f(x_6, x_13, x_14, x_15, x_16, x_36); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_6); +x_44 = l_Lean_Meta_evalNat(x_6, x_13, x_14, x_15, x_16, x_43); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_37); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = lean_box(0); +x_48 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_47, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_46); +x_18 = x_48; +goto block_33; +} +else +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = lean_ctor_get(x_44, 1); +lean_inc(x_49); +lean_dec(x_44); +x_50 = lean_ctor_get(x_45, 0); +lean_inc(x_50); +lean_dec(x_45); +x_51 = lean_nat_dec_le(x_3, x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; +lean_dec(x_50); +lean_dec(x_37); +x_52 = lean_box(0); +x_53 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_52, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_49); +x_18 = x_53; +goto block_33; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_nat_sub(x_50, x_3); +lean_dec(x_50); +x_55 = l_Lean_mkNatLit(x_54); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_56 = l_Lean_Meta_Grind_canon(x_55, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_49); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = l_Lean_Meta_Grind_shareCommon(x_57, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_58); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec(x_59); +x_62 = lean_unsigned_to_nat(1u); +x_63 = lean_nat_add(x_37, x_62); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_60); +x_64 = l_Lean_Meta_Grind_internalize(x_60, x_63, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_61); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +lean_dec(x_64); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_66 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_60, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +if (lean_obj_tag(x_67) == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_37); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_box(0); +x_70 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_69, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_68); +x_18 = x_70; +goto block_33; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_71 = lean_ctor_get(x_66, 1); +lean_inc(x_71); +lean_dec(x_66); +x_72 = lean_ctor_get(x_67, 0); +lean_inc(x_72); +lean_dec(x_67); +x_73 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_72, x_37); +x_74 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_73, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_71); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_75, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_76); +lean_dec(x_75); +x_18 = x_77; +goto block_33; +} +} +else +{ +uint8_t x_78; +lean_dec(x_37); +lean_dec(x_6); +x_78 = !lean_is_exclusive(x_66); +if (x_78 == 0) +{ +x_18 = x_66; +goto block_33; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_66, 0); +x_80 = lean_ctor_get(x_66, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_66); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +x_18 = x_81; +goto block_33; +} +} +} +else +{ +uint8_t x_82; +lean_dec(x_60); +lean_dec(x_37); +lean_dec(x_6); +x_82 = !lean_is_exclusive(x_64); +if (x_82 == 0) +{ +x_18 = x_64; +goto block_33; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_64, 0); +x_84 = lean_ctor_get(x_64, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_64); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_18 = x_85; +goto block_33; +} +} +} +else +{ +uint8_t x_86; +lean_dec(x_37); +lean_dec(x_6); +x_86 = !lean_is_exclusive(x_56); +if (x_86 == 0) +{ +x_18 = x_56; +goto block_33; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_56, 0); +x_88 = lean_ctor_get(x_56, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_56); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_18 = x_89; +goto block_33; +} +} +} +} +} +else +{ +uint8_t x_90; +lean_dec(x_37); +lean_dec(x_6); +x_90 = !lean_is_exclusive(x_44); +if (x_90 == 0) +{ +x_18 = x_44; +goto block_33; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_44, 0); +x_92 = lean_ctor_get(x_44, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_44); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_18 = x_93; +goto block_33; +} +} +} +else +{ +lean_object* x_94; lean_object* x_95; uint8_t x_96; +x_94 = lean_ctor_get(x_42, 0); +lean_inc(x_94); +lean_dec(x_42); +x_95 = lean_ctor_get(x_41, 1); +lean_inc(x_95); +lean_dec(x_41); +x_96 = !lean_is_exclusive(x_94); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_97 = lean_ctor_get(x_94, 0); +x_98 = lean_ctor_get(x_94, 1); +x_99 = lean_nat_dec_lt(x_98, x_3); +if (x_99 == 0) +{ +uint8_t x_100; +lean_free_object(x_94); +x_100 = lean_nat_dec_eq(x_98, x_3); +if (x_100 == 0) +{ +uint8_t x_101; +x_101 = lean_nat_dec_lt(x_3, x_98); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; +lean_dec(x_98); +lean_dec(x_97); +lean_dec(x_37); +x_102 = lean_box(0); +x_103 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_102, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_18 = x_103; +goto block_33; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_104 = lean_nat_sub(x_98, x_3); +lean_dec(x_98); +x_105 = l_Lean_mkNatLit(x_104); +x_106 = l___private_Lean_Expr_0__Lean_natAddFn; +x_107 = l_Lean_mkAppB(x_106, x_97, x_105); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_108 = l_Lean_Meta_Grind_canon(x_107, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +if (lean_obj_tag(x_108) == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = l_Lean_Meta_Grind_shareCommon(x_109, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_110); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = lean_unsigned_to_nat(1u); +x_115 = lean_nat_add(x_37, x_114); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_112); +x_116 = l_Lean_Meta_Grind_internalize(x_112, x_115, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_113); +if (lean_obj_tag(x_116) == 0) +{ +lean_object* x_117; lean_object* x_118; +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +lean_dec(x_116); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_118 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_112, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_117); +if (lean_obj_tag(x_118) == 0) +{ +lean_object* x_119; +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +if (lean_obj_tag(x_119) == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec(x_37); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = lean_box(0); +x_122 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_121, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_120); +x_18 = x_122; +goto block_33; +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_123 = lean_ctor_get(x_118, 1); +lean_inc(x_123); +lean_dec(x_118); +x_124 = lean_ctor_get(x_119, 0); +lean_inc(x_124); +lean_dec(x_119); +x_125 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_124, x_37); +x_126 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_125, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_123); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_127, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_128); +lean_dec(x_127); +x_18 = x_129; +goto block_33; +} +} +else +{ +uint8_t x_130; +lean_dec(x_37); +lean_dec(x_6); +x_130 = !lean_is_exclusive(x_118); +if (x_130 == 0) +{ +x_18 = x_118; +goto block_33; +} +else +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_118, 0); +x_132 = lean_ctor_get(x_118, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_118); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +x_18 = x_133; +goto block_33; +} +} +} +else +{ +uint8_t x_134; +lean_dec(x_112); +lean_dec(x_37); +lean_dec(x_6); +x_134 = !lean_is_exclusive(x_116); +if (x_134 == 0) +{ +x_18 = x_116; +goto block_33; +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_135 = lean_ctor_get(x_116, 0); +x_136 = lean_ctor_get(x_116, 1); +lean_inc(x_136); +lean_inc(x_135); +lean_dec(x_116); +x_137 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +x_18 = x_137; +goto block_33; +} +} +} +else +{ +uint8_t x_138; +lean_dec(x_37); +lean_dec(x_6); +x_138 = !lean_is_exclusive(x_108); +if (x_138 == 0) +{ +x_18 = x_108; +goto block_33; +} +else +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_108, 0); +x_140 = lean_ctor_get(x_108, 1); +lean_inc(x_140); +lean_inc(x_139); +lean_dec(x_108); +x_141 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +x_18 = x_141; +goto block_33; +} +} +} +} +else +{ +lean_object* x_142; +lean_dec(x_98); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_142 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_97, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +if (lean_obj_tag(x_142) == 0) +{ +lean_object* x_143; +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_37); +x_144 = lean_ctor_get(x_142, 1); +lean_inc(x_144); +lean_dec(x_142); +x_145 = lean_box(0); +x_146 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_145, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_144); +x_18 = x_146; +goto block_33; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_147 = lean_ctor_get(x_142, 1); +lean_inc(x_147); +lean_dec(x_142); +x_148 = lean_ctor_get(x_143, 0); +lean_inc(x_148); +lean_dec(x_143); +x_149 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_148, x_37); +x_150 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_149, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_147); +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_153 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_151, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_152); +lean_dec(x_151); +x_18 = x_153; +goto block_33; +} +} +else +{ +uint8_t x_154; +lean_dec(x_37); +lean_dec(x_6); +x_154 = !lean_is_exclusive(x_142); +if (x_154 == 0) +{ +x_18 = x_142; +goto block_33; +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_142, 0); +x_156 = lean_ctor_get(x_142, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_142); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +x_18 = x_157; +goto block_33; +} +} +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; +lean_inc(x_1); +x_158 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_1, x_37); +x_159 = lean_nat_sub(x_3, x_98); +lean_dec(x_98); +lean_inc(x_2); +x_160 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_160, 0, x_2); +lean_ctor_set(x_160, 1, x_159); +lean_ctor_set(x_160, 2, x_97); +x_161 = !lean_is_exclusive(x_158); +if (x_161 == 0) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_162 = lean_ctor_get(x_158, 0); +lean_ctor_set_tag(x_94, 1); +lean_ctor_set(x_94, 1, x_162); +lean_ctor_set(x_94, 0, x_160); +lean_ctor_set(x_158, 0, x_94); +x_163 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_158, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec(x_163); +x_166 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_164, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_165); +lean_dec(x_164); +x_18 = x_166; +goto block_33; +} +else +{ +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_167 = lean_ctor_get(x_158, 0); +x_168 = lean_ctor_get(x_158, 1); +x_169 = lean_ctor_get(x_158, 2); +lean_inc(x_169); +lean_inc(x_168); +lean_inc(x_167); +lean_dec(x_158); +lean_ctor_set_tag(x_94, 1); +lean_ctor_set(x_94, 1, x_167); +lean_ctor_set(x_94, 0, x_160); +x_170 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_170, 0, x_94); +lean_ctor_set(x_170, 1, x_168); +lean_ctor_set(x_170, 2, x_169); +x_171 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_170, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_172, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_173); +lean_dec(x_172); +x_18 = x_174; +goto block_33; +} +} +} +else +{ +lean_object* x_175; lean_object* x_176; uint8_t x_177; +x_175 = lean_ctor_get(x_94, 0); +x_176 = lean_ctor_get(x_94, 1); +lean_inc(x_176); +lean_inc(x_175); +lean_dec(x_94); +x_177 = lean_nat_dec_lt(x_176, x_3); +if (x_177 == 0) +{ +uint8_t x_178; +x_178 = lean_nat_dec_eq(x_176, x_3); +if (x_178 == 0) +{ +uint8_t x_179; +x_179 = lean_nat_dec_lt(x_3, x_176); +if (x_179 == 0) +{ +lean_object* x_180; lean_object* x_181; +lean_dec(x_176); +lean_dec(x_175); +lean_dec(x_37); +x_180 = lean_box(0); +x_181 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_180, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_18 = x_181; +goto block_33; +} +else +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_182 = lean_nat_sub(x_176, x_3); +lean_dec(x_176); +x_183 = l_Lean_mkNatLit(x_182); +x_184 = l___private_Lean_Expr_0__Lean_natAddFn; +x_185 = l_Lean_mkAppB(x_184, x_175, x_183); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_186 = l_Lean_Meta_Grind_canon(x_185, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +if (lean_obj_tag(x_186) == 0) +{ +lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_187 = lean_ctor_get(x_186, 0); +lean_inc(x_187); +x_188 = lean_ctor_get(x_186, 1); +lean_inc(x_188); +lean_dec(x_186); +x_189 = l_Lean_Meta_Grind_shareCommon(x_187, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_188); +x_190 = lean_ctor_get(x_189, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_189, 1); +lean_inc(x_191); +lean_dec(x_189); +x_192 = lean_unsigned_to_nat(1u); +x_193 = lean_nat_add(x_37, x_192); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_190); +x_194 = l_Lean_Meta_Grind_internalize(x_190, x_193, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_191); +if (lean_obj_tag(x_194) == 0) +{ +lean_object* x_195; lean_object* x_196; +x_195 = lean_ctor_get(x_194, 1); +lean_inc(x_195); +lean_dec(x_194); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_196 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_190, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_195); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_197; +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +if (lean_obj_tag(x_197) == 0) +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_37); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +lean_dec(x_196); +x_199 = lean_box(0); +x_200 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_199, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_198); +x_18 = x_200; +goto block_33; +} +else +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_201 = lean_ctor_get(x_196, 1); +lean_inc(x_201); +lean_dec(x_196); +x_202 = lean_ctor_get(x_197, 0); +lean_inc(x_202); +lean_dec(x_197); +x_203 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_202, x_37); +x_204 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_203, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_201); +x_205 = lean_ctor_get(x_204, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_204, 1); +lean_inc(x_206); +lean_dec(x_204); +x_207 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_205, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_206); +lean_dec(x_205); +x_18 = x_207; +goto block_33; +} +} +else +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +lean_dec(x_37); +lean_dec(x_6); +x_208 = lean_ctor_get(x_196, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_196, 1); +lean_inc(x_209); +if (lean_is_exclusive(x_196)) { + lean_ctor_release(x_196, 0); + lean_ctor_release(x_196, 1); + x_210 = x_196; +} else { + lean_dec_ref(x_196); + x_210 = lean_box(0); +} +if (lean_is_scalar(x_210)) { + x_211 = lean_alloc_ctor(1, 2, 0); +} else { + x_211 = x_210; +} +lean_ctor_set(x_211, 0, x_208); +lean_ctor_set(x_211, 1, x_209); +x_18 = x_211; +goto block_33; +} +} +else +{ +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_190); +lean_dec(x_37); +lean_dec(x_6); +x_212 = lean_ctor_get(x_194, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_194, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_214 = x_194; +} else { + lean_dec_ref(x_194); + x_214 = lean_box(0); +} +if (lean_is_scalar(x_214)) { + x_215 = lean_alloc_ctor(1, 2, 0); +} else { + x_215 = x_214; +} +lean_ctor_set(x_215, 0, x_212); +lean_ctor_set(x_215, 1, x_213); +x_18 = x_215; +goto block_33; +} +} +else +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; +lean_dec(x_37); +lean_dec(x_6); +x_216 = lean_ctor_get(x_186, 0); +lean_inc(x_216); +x_217 = lean_ctor_get(x_186, 1); +lean_inc(x_217); +if (lean_is_exclusive(x_186)) { + lean_ctor_release(x_186, 0); + lean_ctor_release(x_186, 1); + x_218 = x_186; +} else { + lean_dec_ref(x_186); + x_218 = lean_box(0); +} +if (lean_is_scalar(x_218)) { + x_219 = lean_alloc_ctor(1, 2, 0); +} else { + x_219 = x_218; +} +lean_ctor_set(x_219, 0, x_216); +lean_ctor_set(x_219, 1, x_217); +x_18 = x_219; +goto block_33; +} +} +} +else +{ +lean_object* x_220; +lean_dec(x_176); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_220 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f(x_1, x_2, x_175, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +if (lean_obj_tag(x_220) == 0) +{ +lean_object* x_221; +x_221 = lean_ctor_get(x_220, 0); +lean_inc(x_221); +if (lean_obj_tag(x_221) == 0) +{ +lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_37); +x_222 = lean_ctor_get(x_220, 1); +lean_inc(x_222); +lean_dec(x_220); +x_223 = lean_box(0); +x_224 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_223, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_222); +x_18 = x_224; +goto block_33; +} +else +{ +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_225 = lean_ctor_get(x_220, 1); +lean_inc(x_225); +lean_dec(x_220); +x_226 = lean_ctor_get(x_221, 0); +lean_inc(x_226); +lean_dec(x_221); +x_227 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_226, x_37); +x_228 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_227, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_225); +x_229 = lean_ctor_get(x_228, 0); +lean_inc(x_229); +x_230 = lean_ctor_get(x_228, 1); +lean_inc(x_230); +lean_dec(x_228); +x_231 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_229, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_230); +lean_dec(x_229); +x_18 = x_231; +goto block_33; +} +} +else +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; +lean_dec(x_37); +lean_dec(x_6); +x_232 = lean_ctor_get(x_220, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_220, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_220)) { + lean_ctor_release(x_220, 0); + lean_ctor_release(x_220, 1); + x_234 = x_220; +} else { + lean_dec_ref(x_220); + x_234 = lean_box(0); +} +if (lean_is_scalar(x_234)) { + x_235 = lean_alloc_ctor(1, 2, 0); +} else { + x_235 = x_234; +} +lean_ctor_set(x_235, 0, x_232); +lean_ctor_set(x_235, 1, x_233); +x_18 = x_235; +goto block_33; +} +} +} +else +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +lean_inc(x_1); +x_236 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_Choice_updateGen(x_1, x_37); +x_237 = lean_nat_sub(x_3, x_176); +lean_dec(x_176); +lean_inc(x_2); +x_238 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_238, 0, x_2); +lean_ctor_set(x_238, 1, x_237); +lean_ctor_set(x_238, 2, x_175); +x_239 = lean_ctor_get(x_236, 0); +lean_inc(x_239); +x_240 = lean_ctor_get(x_236, 1); +lean_inc(x_240); +x_241 = lean_ctor_get(x_236, 2); +lean_inc(x_241); +if (lean_is_exclusive(x_236)) { + lean_ctor_release(x_236, 0); + lean_ctor_release(x_236, 1); + lean_ctor_release(x_236, 2); + x_242 = x_236; +} else { + lean_dec_ref(x_236); + x_242 = lean_box(0); +} +x_243 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_243, 0, x_238); +lean_ctor_set(x_243, 1, x_239); +if (lean_is_scalar(x_242)) { + x_244 = lean_alloc_ctor(0, 3, 0); +} else { + x_244 = x_242; +} +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_240); +lean_ctor_set(x_244, 2, x_241); +x_245 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_pushChoice(x_244, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); +x_246 = lean_ctor_get(x_245, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_245, 1); +lean_inc(x_247); +lean_dec(x_245); +x_248 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch___spec__1___lambda__1(x_4, x_6, x_246, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_247); +lean_dec(x_246); +x_18 = x_248; +goto block_33; +} +} +} +} +else +{ +uint8_t x_249; +lean_dec(x_37); +lean_dec(x_6); +x_249 = !lean_is_exclusive(x_41); +if (x_249 == 0) +{ +x_18 = x_41; +goto block_33; +} +else +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_250 = lean_ctor_get(x_41, 0); +x_251 = lean_ctor_get(x_41, 1); +lean_inc(x_251); +lean_inc(x_250); +lean_dec(x_41); +x_252 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +x_18 = x_252; +goto block_33; +} +} +} +} +else +{ +uint8_t x_253; +lean_dec(x_6); +x_253 = !lean_is_exclusive(x_34); +if (x_253 == 0) +{ +x_18 = x_34; +goto block_33; +} +else +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_254 = lean_ctor_get(x_34, 0); +x_255 = lean_ctor_get(x_34, 1); +lean_inc(x_255); +lean_inc(x_254); +lean_dec(x_34); +x_256 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +x_18 = x_256; +goto block_33; +} +} +block_33: +{ +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 0); +lean_dec(x_21); +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +lean_ctor_set(x_18, 0, x_22); +return x_18; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_dec(x_18); +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; +} +} +else +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_18, 1); +lean_inc(x_26); +lean_dec(x_18); +x_27 = lean_ctor_get(x_19, 0); +lean_inc(x_27); +lean_dec(x_19); +x_6 = x_27; +x_17 = x_26; +goto _start; +} +} +else +{ +uint8_t x_29; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_18); +if (x_29 == 0) +{ +return x_18; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_18, 0); +x_31 = lean_ctor_get(x_18, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_18); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_4); +x_19 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1(x_1, x_2, x_3, x_4, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +lean_dec(x_17); +lean_dec(x_4); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +x_22 = lean_box(0); +lean_ctor_set(x_19, 0, x_22); +return x_19; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_19, 1); +lean_inc(x_23); +lean_dec(x_19); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) +{ +return x_19; +} +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_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_16; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +_start: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_21; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_30; +lean_dec(x_8); +x_22 = lean_ctor_get(x_7, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_7, 1); +lean_inc(x_23); +if (lean_is_exclusive(x_7)) { + lean_ctor_release(x_7, 0); + lean_ctor_release(x_7, 1); + x_24 = x_7; +} else { + lean_dec_ref(x_7); + x_24 = lean_box(0); +} +lean_inc(x_22); +x_30 = l_Lean_Meta_Grind_getENode(x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_83; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 8); +lean_inc(x_33); +x_83 = lean_nat_dec_lt(x_33, x_4); +if (x_83 == 0) +{ +lean_object* x_84; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_24); +lean_dec(x_22); +x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_84; +x_26 = x_32; +goto block_29; +} +else +{ +uint8_t x_85; +x_85 = lean_ctor_get_uint8(x_31, sizeof(void*)*10 + 4); +if (x_85 == 0) +{ +uint8_t x_86; +x_86 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_31); +lean_dec(x_31); +if (x_86 == 0) +{ +lean_object* x_87; +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_22); +x_87 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_87; +x_26 = x_32; +goto block_29; +} +else +{ +lean_object* x_88; +x_88 = lean_box(0); +x_34 = x_88; +goto block_82; +} +} +else +{ +lean_object* x_89; +lean_dec(x_31); +x_89 = lean_box(0); +x_34 = x_89; +goto block_82; +} +} +block_82: +{ +lean_object* x_35; +lean_dec(x_34); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_1); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_1, x_2, x_22, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_32); +lean_dec(x_22); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_33); +lean_dec(x_24); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_38; +x_26 = x_37; +goto block_29; +} +else +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_dec(x_35); +x_41 = !lean_is_exclusive(x_39); +if (x_41 == 0) +{ +lean_object* x_42; uint8_t x_43; +x_42 = lean_ctor_get(x_39, 1); +x_43 = lean_nat_dec_le(x_33, x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_42); +lean_ctor_set(x_39, 1, x_33); +x_44 = lean_st_ref_take(x_11, x_40); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +if (lean_is_scalar(x_24)) { + x_47 = lean_alloc_ctor(1, 2, 0); +} else { + x_47 = x_24; +} +lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_47, 1, x_45); +x_48 = lean_st_ref_set(x_11, x_47, x_46); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_50; +x_26 = x_49; +goto block_29; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_33); +x_51 = lean_st_ref_take(x_11, x_40); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +if (lean_is_scalar(x_24)) { + x_54 = lean_alloc_ctor(1, 2, 0); +} else { + x_54 = x_24; +} +lean_ctor_set(x_54, 0, x_39); +lean_ctor_set(x_54, 1, x_52); +x_55 = lean_st_ref_set(x_11, x_54, x_53); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_57; +x_26 = x_56; +goto block_29; +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_58 = lean_ctor_get(x_39, 0); +x_59 = lean_ctor_get(x_39, 1); +x_60 = lean_ctor_get(x_39, 2); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_39); +x_61 = lean_nat_dec_le(x_33, x_59); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_59); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_58); +lean_ctor_set(x_62, 1, x_33); +lean_ctor_set(x_62, 2, x_60); +x_63 = lean_st_ref_take(x_11, x_40); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +if (lean_is_scalar(x_24)) { + x_66 = lean_alloc_ctor(1, 2, 0); +} else { + x_66 = x_24; +} +lean_ctor_set(x_66, 0, x_62); +lean_ctor_set(x_66, 1, x_64); +x_67 = lean_st_ref_set(x_11, x_66, x_65); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +lean_dec(x_67); +x_69 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_69; +x_26 = x_68; +goto block_29; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_33); +x_70 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_70, 0, x_58); +lean_ctor_set(x_70, 1, x_59); +lean_ctor_set(x_70, 2, x_60); +x_71 = lean_st_ref_take(x_11, x_40); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +if (lean_is_scalar(x_24)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_24; +} +lean_ctor_set(x_74, 0, x_70); +lean_ctor_set(x_74, 1, x_72); +x_75 = lean_st_ref_set(x_11, x_74, x_73); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1; +x_25 = x_77; +x_26 = x_76; +goto block_29; +} +} +} +} +else +{ +uint8_t x_78; +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_78 = !lean_is_exclusive(x_35); +if (x_78 == 0) +{ +return x_35; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_35, 0); +x_80 = lean_ctor_get(x_35, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_35); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +return x_81; +} +} +} +} +else +{ +uint8_t x_90; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_1); +x_90 = !lean_is_exclusive(x_30); +if (x_90 == 0) +{ +return x_30; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_30, 0); +x_92 = lean_ctor_get(x_30, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_30); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; +} +} +block_29: +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_7 = x_23; +x_8 = x_27; +x_9 = lean_box(0); +x_20 = x_26; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_st_ref_get(x_5, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +x_18 = lean_ctor_get(x_16, 4); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_2); +x_19 = l_Lean_Expr_toHeadIndex(x_2); +x_20 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_18, x_19); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +lean_ctor_set(x_14, 0, x_21); +return x_14; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_free_object(x_14); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_17); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_box(0); +x_27 = lean_box(0); +lean_inc(x_22); +x_28 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_22, x_24, x_26, x_22, x_22, x_27, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_25); +lean_dec(x_24); +lean_dec(x_22); +lean_dec(x_2); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +lean_object* x_30; +x_30 = lean_ctor_get(x_28, 0); +lean_dec(x_30); +lean_ctor_set(x_28, 0, x_27); +return x_28; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_27); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) +{ +return x_28; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_14, 0); +x_38 = lean_ctor_get(x_14, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_14); +x_39 = lean_ctor_get(x_37, 4); +lean_inc(x_39); +lean_dec(x_37); +lean_inc(x_2); +x_40 = l_Lean_Expr_toHeadIndex(x_2); +x_41 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_39, x_40); +lean_dec(x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_38); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_44 = lean_ctor_get(x_41, 0); +lean_inc(x_44); +lean_dec(x_41); +x_45 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_38); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_box(0); +x_49 = lean_box(0); +lean_inc(x_44); +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_44, x_46, x_48, x_44, x_44, x_49, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); +lean_dec(x_46); +lean_dec(x_44); +lean_dec(x_2); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_52 = x_50; +} else { + lean_dec_ref(x_50); + x_52 = lean_box(0); +} +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_52; +} +lean_ctor_set(x_53, 0, x_49); +lean_ctor_set(x_53, 1, x_51); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_50, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_50, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_56 = x_50; +} else { + lean_dec_ref(x_50); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(1, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +_start: +{ +lean_object* x_21; +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +lean_dec(x_3); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = lean_apply_12(x_1, x_8, x_2, x_3, x_4, x_5, x_6, x_7, x_9, x_10, x_11, x_12, x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___lambda__1), 13, 7); +lean_closure_set(x_17, 0, x_4); +lean_closure_set(x_17, 1, x_6); +lean_closure_set(x_17, 2, x_7); +lean_closure_set(x_17, 3, x_8); +lean_closure_set(x_17, 4, x_9); +lean_closure_set(x_17, 5, x_10); +lean_closure_set(x_17, 6, x_11); +x_18 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_17, x_5, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +return x_18; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) +{ +return x_18; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_18, 0); +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_18); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___boxed), 16, 0); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("EqMatch", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Grind_markAsDoNotSimp(x_4, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_17, 0); +x_20 = l_Lean_Expr_constLevels_x21(x_2); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4; +x_22 = l_Lean_Expr_const___override(x_21, x_20); +x_23 = l_Lean_mkApp4(x_22, x_3, x_19, x_5, x_1); +lean_ctor_set(x_17, 0, x_23); +return x_17; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_17, 0); +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_17); +x_26 = l_Lean_Expr_constLevels_x21(x_2); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4; +x_28 = l_Lean_Expr_const___override(x_27, x_26); +x_29 = l_Lean_mkApp4(x_28, x_3, x_24, x_5, x_1); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +return x_30; +} +} +else +{ +uint8_t x_31; +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_17); +if (x_31 == 0) +{ +return x_17; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_17, 0); +x_33 = lean_ctor_get(x_17, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_17); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_expr_instantiate1(x_1, x_3); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType(x_15, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_3); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_array_mk(x_20); +x_22 = 0; +x_23 = 1; +x_24 = 1; +x_25 = l_Lean_Meta_mkForallFVars(x_21, x_17, x_22, x_23, x_24, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_21); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +x_26 = !lean_is_exclusive(x_16); +if (x_26 == 0) +{ +return x_16; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_16, 0); +x_28 = lean_ctor_get(x_16, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_16); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +if (lean_obj_tag(x_1) == 7) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 2); +lean_inc(x_16); +x_17 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_dec(x_1); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_18 = l_Lean_Meta_Grind_markAsDoNotSimp(x_15, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3___boxed), 14, 2); +lean_closure_set(x_21, 0, x_16); +lean_closure_set(x_21, 1, x_2); +x_22 = 0; +x_23 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg(x_14, x_17, x_19, x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_20); +return x_23; +} +else +{ +uint8_t x_24; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +return x_18; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_18, 0); +x_26 = lean_ctor_get(x_18, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_18); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +else +{ +lean_object* x_28; uint8_t x_29; +lean_inc(x_1); +x_28 = l_Lean_Expr_cleanupAnnotations(x_1); +x_29 = l_Lean_Expr_isApp(x_28); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_28); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_13); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = l_Lean_Expr_appArg(x_28, lean_box(0)); +x_32 = l_Lean_Expr_appFnCleanup(x_28, lean_box(0)); +x_33 = l_Lean_Expr_isApp(x_32); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_1); +lean_ctor_set(x_34, 1, x_13); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Expr_appArg(x_32, lean_box(0)); +x_36 = l_Lean_Expr_appFnCleanup(x_32, lean_box(0)); +x_37 = l_Lean_Expr_isApp(x_36); +if (x_37 == 0) +{ +lean_object* x_38; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_31); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_1); +lean_ctor_set(x_38, 1, x_13); +return x_38; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_39 = l_Lean_Expr_appArg(x_36, lean_box(0)); +x_40 = l_Lean_Expr_appFnCleanup(x_36, lean_box(0)); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2; +x_42 = l_Lean_Expr_isConstOf(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_35); +lean_dec(x_31); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_1); +lean_ctor_set(x_43, 1, x_13); +return x_43; +} +else +{ +lean_object* x_44; +lean_dec(x_1); +x_44 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1(x_2, x_40, x_39, x_35, x_31, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_40); +return x_44; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_unbox(x_2); +lean_dec(x_2); +x_18 = lean_unbox(x_5); +lean_dec(x_5); +x_19 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___spec__1___rarg(x_1, x_17, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_19; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = l_Lean_Expr_hasMVar(x_1); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +else { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; x_15 = lean_st_ref_get(x_9, x_12); @@ -2958,8 +5475,7 @@ x_16 = lean_ctor_get(x_13, 1); x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -lean_inc(x_1); -x_18 = lean_environment_find(x_17, x_1); +x_18 = l_Lean_Environment_find_x3f(x_17, x_1); if (lean_obj_tag(x_18) == 0) { uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; @@ -2999,8 +5515,7 @@ lean_dec(x_13); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -lean_inc(x_1); -x_30 = lean_environment_find(x_29, x_1); +x_30 = l_Lean_Environment_find_x3f(x_29, x_1); if (lean_obj_tag(x_30) == 0) { uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; @@ -3098,33 +5613,6 @@ return x_30; } } } -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("[unknown]", 9, 9); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -3230,12 +5718,15 @@ return x_36; } default: { -lean_object* x_37; lean_object* x_38; -x_37 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3; -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_12); -return x_38; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_37); +lean_dec(x_1); +x_38 = l_Lean_MessageData_ofName(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_12); +return x_39; } } } @@ -3565,79 +6056,313 @@ if (lean_is_exclusive(x_103)) { lean_dec_ref(x_103); x_105 = lean_box(0); } -x_106 = lean_box(0); -if (lean_is_scalar(x_105)) { - x_107 = lean_alloc_ctor(0, 2, 0); -} else { - x_107 = x_105; +x_106 = lean_box(0); +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_105; +} +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_104); +return x_107; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_add(x_1, x_16); +x_18 = l_Lean_Meta_Grind_addTheoremInstance(x_2, x_3, x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_18; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematch", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("instance", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4; +x_18 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_3); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_4, x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_21); +return x_23; +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_18, 1); +x_26 = lean_ctor_get(x_18, 0); +lean_dec(x_26); +x_27 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_25); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_28); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_18, 7); +lean_ctor_set(x_18, 1, x_30); +lean_ctor_set(x_18, 0, x_32); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_18); +lean_ctor_set(x_34, 1, x_33); +lean_inc(x_4); +x_35 = l_Lean_MessageData_ofExpr(x_4); +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_32); +x_38 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_17, x_37, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_31); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_4, x_39, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_40); +lean_dec(x_39); +return x_41; +} +else +{ +uint8_t x_42; +lean_free_object(x_18); +lean_dec(x_4); +lean_dec(x_2); +x_42 = !lean_is_exclusive(x_29); +if (x_42 == 0) +{ +return x_29; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_29, 0); +x_44 = lean_ctor_get(x_29, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_29); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_104); -return x_107; } } +else +{ +uint8_t x_46; +lean_free_object(x_18); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_46 = !lean_is_exclusive(x_27); +if (x_46 == 0) +{ +return x_27; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +else { -lean_object* x_16; -x_16 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_2, x_3, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -return x_16; +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_27, 0); +x_48 = lean_ctor_get(x_27, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_27); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; } +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_18, 1); +lean_inc(x_50); +lean_dec(x_18); +x_51 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_50); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_52); +if (lean_obj_tag(x_53) == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +x_58 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +lean_inc(x_4); +x_60 = l_Lean_MessageData_ofExpr(x_4); +x_61 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_56); +x_63 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_17, x_62, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_55); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__1(x_1, x_2, x_4, x_64, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_65); +lean_dec(x_64); +return x_66; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ematch", 6, 6); -return x_1; +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_4); +lean_dec(x_2); +x_67 = lean_ctor_get(x_53, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_53, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_69 = x_53; +} else { + lean_dec_ref(x_53); + x_69 = lean_box(0); } +if (lean_is_scalar(x_69)) { + x_70 = lean_alloc_ctor(1, 2, 0); +} else { + x_70 = x_69; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("instance", 8, 8); -return x_1; +lean_ctor_set(x_70, 0, x_67); +lean_ctor_set(x_70, 1, x_68); +return x_70; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_71 = lean_ctor_get(x_51, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_51, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_73 = x_51; +} else { + lean_dec_ref(x_51); + x_73 = lean_box(0); } +if (lean_is_scalar(x_73)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_73; +} +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +return x_74; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); -return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; @@ -3649,212 +6374,148 @@ lean_inc(x_1); x_16 = lean_infer_type(x_1, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__4; -x_20 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); -x_21 = lean_ctor_get(x_20, 0); +x_19 = lean_st_ref_get(x_14, x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); -x_22 = lean_unbox(x_21); -lean_dec(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; -lean_dec(x_3); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); +lean_dec(x_19); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); lean_dec(x_20); -x_24 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_23); +x_23 = l_Lean_Meta_Grind_Origin_key(x_3); +x_24 = l_Lean_Meta_Match_isMatchEqnTheorem(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_box(0); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_2, x_1, x_3, x_17, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_21); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -return x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_26; } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_20); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_20, 1); -x_27 = lean_ctor_get(x_20, 0); -lean_dec(x_27); -x_28 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_26); +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_5, 1); +lean_inc(x_27); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType(x_17, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_21); if (lean_obj_tag(x_28) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_29); -lean_ctor_set(x_20, 0, x_31); -x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_20); -lean_ctor_set(x_33, 1, x_32); -lean_inc(x_17); -x_34 = l_Lean_MessageData_ofExpr(x_17); -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_31); -x_37 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_19, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_30); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_38); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -return x_39; -} -else -{ -uint8_t x_40; -lean_free_object(x_20); -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_2); -lean_dec(x_1); -x_40 = !lean_is_exclusive(x_28); -if (x_40 == 0) -{ -return x_28; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_28, 0); -x_42 = lean_ctor_get(x_28, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_28); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; -} -} -} -else -{ -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_20, 1); -lean_inc(x_44); -lean_dec(x_20); -x_45 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_44); -if (lean_obj_tag(x_45) == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -x_48 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_46); -x_50 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; -x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -lean_inc(x_17); -x_52 = l_Lean_MessageData_ofExpr(x_17); -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_48); -x_55 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_19, x_54, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_47); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_17, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_56); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_2, x_1, x_3, x_29, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_30); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -return x_57; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_32; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_dec(x_17); +uint8_t x_33; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_2); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_1); -x_58 = lean_ctor_get(x_45, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_45, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_45)) { - lean_ctor_release(x_45, 0); - lean_ctor_release(x_45, 1); - x_60 = x_45; -} else { - lean_dec_ref(x_45); - x_60 = lean_box(0); -} -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(1, 2, 0); -} else { - x_61 = x_60; +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) +{ +return x_28; } -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_59); -return x_61; +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } } else { -uint8_t x_62; +uint8_t x_37; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_62 = !lean_is_exclusive(x_16); -if (x_62 == 0) +x_37 = !lean_is_exclusive(x_16); +if (x_37 == 0) { return x_16; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_16, 0); -x_64 = lean_ctor_get(x_16, 1); -lean_inc(x_64); -lean_inc(x_63); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_16, 0); +x_39 = lean_ctor_get(x_16, 1); +lean_inc(x_39); +lean_inc(x_38); lean_dec(x_16); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +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; } } } @@ -3886,7 +6547,7 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_16, x_3, x_1, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3(x_16, x_3, x_1, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_17); return x_22; } else @@ -3906,7 +6567,7 @@ lean_inc(x_24); x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); -x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_16, x_3, x_1, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_25); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3(x_16, x_3, x_1, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_25); lean_dec(x_24); return x_26; } @@ -3918,7 +6579,12 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_3); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_1); x_27 = !lean_is_exclusive(x_23); if (x_27 == 0) @@ -4084,21 +6750,37 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_1); return x_16; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_16; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_4); +lean_dec(x_2); return x_16; } } @@ -4107,12 +6789,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_ { lean_object* x_15; x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_3); return x_15; } } @@ -4376,6 +7053,23 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23) { _start: { @@ -4405,29 +7099,29 @@ return x_26; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_36; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; +lean_object* x_27; lean_object* x_28; lean_object* x_36; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; lean_dec(x_9); -x_165 = lean_nat_sub(x_3, x_10); -x_166 = lean_unsigned_to_nat(1u); -x_167 = lean_nat_sub(x_165, x_166); -lean_dec(x_165); -x_168 = lean_nat_dec_lt(x_167, x_4); -if (x_168 == 0) +x_197 = lean_nat_sub(x_3, x_10); +x_198 = lean_unsigned_to_nat(1u); +x_199 = lean_nat_sub(x_197, x_198); +lean_dec(x_197); +x_200 = lean_nat_dec_lt(x_199, x_4); +if (x_200 == 0) { -lean_object* x_169; lean_object* x_170; -lean_dec(x_167); -x_169 = l_Lean_instInhabitedExpr; -x_170 = l_outOfBounds___rarg(x_169); -x_36 = x_170; -goto block_164; +lean_object* x_201; lean_object* x_202; +lean_dec(x_199); +x_201 = l_Lean_instInhabitedExpr; +x_202 = l_outOfBounds___rarg(x_201); +x_36 = x_202; +goto block_196; } else { -lean_object* x_171; -x_171 = lean_array_fget(x_2, x_167); -lean_dec(x_167); -x_36 = x_171; -goto block_164; +lean_object* x_203; +x_203 = lean_array_fget(x_2, x_199); +lean_dec(x_199); +x_36 = x_203; +goto block_196; } block_35: { @@ -4472,96 +7166,355 @@ x_23 = x_28; goto _start; } } -block_164: +block_196: { lean_object* x_37; uint8_t x_38; x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; x_38 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_36, x_37); if (x_38 == 0) { -lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_130; lean_object* x_131; +lean_object* x_39; lean_object* x_40; lean_object* x_41; x_39 = lean_array_fget(x_5, x_10); -x_130 = l_Lean_Expr_mvarId_x21(x_39); -lean_inc(x_130); -x_131 = l_Lean_MVarId_getType(x_130, x_19, x_20, x_21, x_22, x_23); -if (lean_obj_tag(x_131) == 0) +x_40 = l_Lean_Expr_mvarId_x21(x_39); +lean_inc(x_40); +x_41 = l_Lean_MVarId_getType(x_40, x_19, x_20, x_21, x_22, x_23); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); lean_inc(x_36); -x_134 = lean_infer_type(x_36, x_19, x_20, x_21, x_22, x_133); -if (lean_obj_tag(x_134) == 0) +x_44 = lean_infer_type(x_36, x_19, x_20, x_21, x_22, x_43); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -lean_dec(x_134); +lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_169; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -x_137 = l_Lean_Meta_isExprDefEq(x_132, x_135, x_19, x_20, x_21, x_22, x_136); -if (lean_obj_tag(x_137) == 0) +lean_inc(x_45); +lean_inc(x_42); +x_169 = l_Lean_Meta_isExprDefEq(x_42, x_45, x_19, x_20, x_21, x_22, x_46); +if (lean_obj_tag(x_169) == 0) { -lean_object* x_138; uint8_t x_139; -x_138 = lean_ctor_get(x_137, 0); -lean_inc(x_138); -x_139 = lean_unbox(x_138); -if (x_139 == 0) +lean_object* x_170; uint8_t x_171; +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_unbox(x_170); +if (x_171 == 0) { -lean_object* x_140; uint8_t x_141; -lean_dec(x_130); -x_140 = lean_ctor_get(x_137, 1); -lean_inc(x_140); -lean_dec(x_137); -x_141 = lean_unbox(x_138); -lean_dec(x_138); -x_40 = x_141; -x_41 = x_140; -goto block_129; +lean_object* x_172; uint8_t x_173; +lean_dec(x_40); +x_172 = lean_ctor_get(x_169, 1); +lean_inc(x_172); +lean_dec(x_169); +x_173 = lean_unbox(x_170); +lean_dec(x_170); +x_47 = x_173; +x_48 = x_172; +goto block_168; } else { -lean_object* x_142; lean_object* x_143; -lean_dec(x_138); -x_142 = lean_ctor_get(x_137, 1); -lean_inc(x_142); -lean_dec(x_137); +lean_object* x_174; lean_object* x_175; +lean_dec(x_170); +x_174 = lean_ctor_get(x_169, 1); +lean_inc(x_174); +lean_dec(x_169); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_36); +x_175 = lean_checked_assign(x_40, x_36, x_19, x_20, x_21, x_22, x_174); +if (lean_obj_tag(x_175) == 0) +{ +lean_object* x_176; lean_object* x_177; uint8_t x_178; +x_176 = lean_ctor_get(x_175, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_175, 1); +lean_inc(x_177); +lean_dec(x_175); +x_178 = lean_unbox(x_176); +lean_dec(x_176); +x_47 = x_178; +x_48 = x_177; +goto block_168; +} +else +{ +uint8_t x_179; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_179 = !lean_is_exclusive(x_175); +if (x_179 == 0) +{ +return x_175; +} +else +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_175, 0); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +lean_inc(x_180); +lean_dec(x_175); +x_182 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +return x_182; +} +} +} +} +else +{ +uint8_t x_183; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_183 = !lean_is_exclusive(x_169); +if (x_183 == 0) +{ +return x_169; +} +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_184 = lean_ctor_get(x_169, 0); +x_185 = lean_ctor_get(x_169, 1); +lean_inc(x_185); +lean_inc(x_184); +lean_dec(x_169); +x_186 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +return x_186; +} +} +block_168: +{ +if (x_47 == 0) +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_50 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_49, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_48); +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_50, 1); +x_54 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; +x_55 = lean_unbox(x_52); +lean_dec(x_52); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +lean_free_object(x_50); +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); +x_56 = lean_box(0); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_57 = lean_apply_12(x_54, x_56, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_53); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_27 = x_58; +x_28 = x_59; +goto block_35; +} +else +{ +uint8_t x_60; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_60 = !lean_is_exclusive(x_57); +if (x_60 == 0) +{ +return x_57; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_57, 0); +x_62 = lean_ctor_get(x_57, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_57); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; +} +} +} +else +{ +lean_object* x_64; +x_64 = l_Lean_Meta_Grind_updateLastTag(x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_53); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +lean_dec(x_64); +x_66 = lean_ctor_get(x_1, 5); +lean_inc(x_66); +x_67 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_66, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_65); +if (lean_obj_tag(x_67) == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +x_70 = l_Lean_Meta_mkHasTypeButIsExpectedMsg(x_45, x_42, x_19, x_20, x_21, x_22, x_69); +if (lean_obj_tag(x_70) == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); +x_73 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; +lean_ctor_set_tag(x_50, 7); +lean_ctor_set(x_50, 1, x_68); +lean_ctor_set(x_50, 0, x_73); +x_74 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_50); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Lean_MessageData_ofExpr(x_39); +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +x_78 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; +x_79 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Lean_indentExpr(x_36); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +x_82 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11; +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_71); +x_85 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_86 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_49, x_86, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_72); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); -lean_inc(x_36); -x_143 = lean_checked_assign(x_130, x_36, x_19, x_20, x_21, x_22, x_142); -if (lean_obj_tag(x_143) == 0) +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_90 = lean_apply_12(x_54, x_88, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_89); +if (lean_obj_tag(x_90) == 0) { -lean_object* x_144; lean_object* x_145; uint8_t x_146; -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_143, 1); -lean_inc(x_145); -lean_dec(x_143); -x_146 = lean_unbox(x_144); -lean_dec(x_144); -x_40 = x_146; -x_41 = x_145; -goto block_129; +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_27 = x_91; +x_28 = x_92; +goto block_35; } else { -uint8_t x_147; -lean_dec(x_39); -lean_dec(x_36); +uint8_t x_93; lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -4575,31 +7528,31 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_147 = !lean_is_exclusive(x_143); -if (x_147 == 0) +x_93 = !lean_is_exclusive(x_90); +if (x_93 == 0) { -return x_143; +return x_90; } else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_148 = lean_ctor_get(x_143, 0); -x_149 = lean_ctor_get(x_143, 1); -lean_inc(x_149); -lean_inc(x_148); -lean_dec(x_143); -x_150 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_150, 0, x_148); -lean_ctor_set(x_150, 1, x_149); -return x_150; -} +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_90, 0); +x_95 = lean_ctor_get(x_90, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_90); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } else { -uint8_t x_151; -lean_dec(x_130); +uint8_t x_97; +lean_dec(x_68); +lean_free_object(x_50); lean_dec(x_39); lean_dec(x_36); lean_dec(x_22); @@ -4615,31 +7568,32 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_151 = !lean_is_exclusive(x_137); -if (x_151 == 0) +x_97 = !lean_is_exclusive(x_70); +if (x_97 == 0) { -return x_137; +return x_70; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = lean_ctor_get(x_137, 0); -x_153 = lean_ctor_get(x_137, 1); -lean_inc(x_153); -lean_inc(x_152); -lean_dec(x_137); -x_154 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_154, 0, x_152); -lean_ctor_set(x_154, 1, x_153); -return x_154; +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_70, 0); +x_99 = lean_ctor_get(x_70, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_70); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; } } } else { -uint8_t x_155; -lean_dec(x_132); -lean_dec(x_130); +uint8_t x_101; +lean_free_object(x_50); +lean_dec(x_45); +lean_dec(x_42); lean_dec(x_39); lean_dec(x_36); lean_dec(x_22); @@ -4655,30 +7609,32 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_155 = !lean_is_exclusive(x_134); -if (x_155 == 0) +x_101 = !lean_is_exclusive(x_67); +if (x_101 == 0) { -return x_134; +return x_67; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_134, 0); -x_157 = lean_ctor_get(x_134, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_134); -x_158 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -return x_158; +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_67, 0); +x_103 = lean_ctor_get(x_67, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_67); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; } } } else { -uint8_t x_159; -lean_dec(x_130); +uint8_t x_105; +lean_free_object(x_50); +lean_dec(x_45); +lean_dec(x_42); lean_dec(x_39); lean_dec(x_36); lean_dec(x_22); @@ -4694,48 +7650,46 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_159 = !lean_is_exclusive(x_131); -if (x_159 == 0) +x_105 = !lean_is_exclusive(x_64); +if (x_105 == 0) { -return x_131; +return x_64; } else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_160 = lean_ctor_get(x_131, 0); -x_161 = lean_ctor_get(x_131, 1); -lean_inc(x_161); -lean_inc(x_160); -lean_dec(x_131); -x_162 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_162, 0, x_160); -lean_ctor_set(x_162, 1, x_161); -return x_162; +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_64, 0); +x_107 = lean_ctor_get(x_64, 1); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_64); +x_108 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +return x_108; } } -block_129: -{ -if (x_40 == 0) +} +} +else { -lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_42 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_43 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_42, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_41); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_109 = lean_ctor_get(x_50, 0); +x_110 = lean_ctor_get(x_50, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_50); +x_111 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; +x_112 = lean_unbox(x_109); +lean_dec(x_109); +if (x_112 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_43, 0); -x_46 = lean_ctor_get(x_43, 1); -x_47 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; -x_48 = lean_unbox(x_45); +lean_object* x_113; lean_object* x_114; lean_dec(x_45); -if (x_48 == 0) -{ -lean_object* x_49; lean_object* x_50; -lean_free_object(x_43); +lean_dec(x_42); lean_dec(x_39); lean_dec(x_36); -x_49 = lean_box(0); +x_113 = lean_box(0); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); @@ -4746,22 +7700,22 @@ lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); -x_50 = lean_apply_12(x_47, x_49, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_46); -if (lean_obj_tag(x_50) == 0) +x_114 = lean_apply_12(x_111, x_113, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_110); +if (lean_obj_tag(x_114) == 0) { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); -x_27 = x_51; -x_28 = x_52; +lean_object* x_115; lean_object* x_116; +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_114, 1); +lean_inc(x_116); +lean_dec(x_114); +x_27 = x_115; +x_28 = x_116; goto block_35; } else { -uint8_t x_53; +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -4775,70 +7729,99 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_53 = !lean_is_exclusive(x_50); -if (x_53 == 0) -{ -return x_50; +x_117 = lean_ctor_get(x_114, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_114, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_119 = x_114; +} else { + lean_dec_ref(x_114); + x_119 = lean_box(0); } -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_50, 0); -x_55 = lean_ctor_get(x_50, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_50); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 2, 0); +} else { + x_120 = x_119; } +lean_ctor_set(x_120, 0, x_117); +lean_ctor_set(x_120, 1, x_118); +return x_120; } } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_1, 5); -lean_inc(x_57); -x_58 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_57, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_46); -if (lean_obj_tag(x_58) == 0) +lean_object* x_121; +x_121 = l_Lean_Meta_Grind_updateLastTag(x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_110); +if (lean_obj_tag(x_121) == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -x_61 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; -lean_ctor_set_tag(x_43, 7); -lean_ctor_set(x_43, 1, x_59); -lean_ctor_set(x_43, 0, x_61); -x_62 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_43); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_MessageData_ofExpr(x_39); -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; -x_67 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -x_68 = l_Lean_indentExpr(x_36); -x_69 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -x_70 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_71 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_42, x_71, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_60); -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); -lean_inc(x_74); -lean_dec(x_72); +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_121, 1); +lean_inc(x_122); +lean_dec(x_121); +x_123 = lean_ctor_get(x_1, 5); +lean_inc(x_123); +x_124 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_123, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_122); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +x_127 = l_Lean_Meta_mkHasTypeButIsExpectedMsg(x_45, x_42, x_19, x_20, x_21, x_22, x_126); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +lean_dec(x_127); +x_130 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; +x_131 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_125); +x_132 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; +x_133 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +x_134 = l_Lean_MessageData_ofExpr(x_39); +x_135 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +x_136 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; +x_137 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +x_138 = l_Lean_indentExpr(x_36); +x_139 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_139, 0, x_137); +lean_ctor_set(x_139, 1, x_138); +x_140 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11; +x_141 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +x_142 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_128); +x_143 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_144 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +x_145 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_49, x_144, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_129); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +lean_dec(x_145); lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); @@ -4849,61 +7832,22 @@ lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); -x_75 = lean_apply_12(x_47, x_73, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_74); -if (lean_obj_tag(x_75) == 0) +x_148 = lean_apply_12(x_111, x_146, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_147); +if (lean_obj_tag(x_148) == 0) { -lean_object* x_76; lean_object* x_77; -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -lean_dec(x_75); -x_27 = x_76; -x_28 = x_77; +lean_object* x_149; lean_object* x_150; +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +lean_dec(x_148); +x_27 = x_149; +x_28 = x_150; goto block_35; } else { -uint8_t x_78; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_78 = !lean_is_exclusive(x_75); -if (x_78 == 0) -{ -return x_75; -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_75, 0); -x_80 = lean_ctor_get(x_75, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_75); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -return x_81; -} -} -} -else -{ -uint8_t x_82; -lean_free_object(x_43); -lean_dec(x_39); -lean_dec(x_36); +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -4917,70 +7861,34 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_82 = !lean_is_exclusive(x_58); -if (x_82 == 0) -{ -return x_58; -} -else -{ -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_58, 0); -x_84 = lean_ctor_get(x_58, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_58); -x_85 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -return x_85; -} -} -} +x_151 = lean_ctor_get(x_148, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_148, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_153 = x_148; +} else { + lean_dec_ref(x_148); + x_153 = lean_box(0); +} +if (lean_is_scalar(x_153)) { + x_154 = lean_alloc_ctor(1, 2, 0); +} else { + x_154 = x_153; +} +lean_ctor_set(x_154, 0, x_151); +lean_ctor_set(x_154, 1, x_152); +return x_154; } -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_86 = lean_ctor_get(x_43, 0); -x_87 = lean_ctor_get(x_43, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_43); -x_88 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__3; -x_89 = lean_unbox(x_86); -lean_dec(x_86); -if (x_89 == 0) -{ -lean_object* x_90; lean_object* x_91; -lean_dec(x_39); -lean_dec(x_36); -x_90 = lean_box(0); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_91 = lean_apply_12(x_88, x_90, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_87); -if (lean_obj_tag(x_91) == 0) -{ -lean_object* x_92; lean_object* x_93; -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -lean_dec(x_91); -x_27 = x_92; -x_28 = x_93; -goto block_35; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_125); +lean_dec(x_39); +lean_dec(x_36); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -4994,98 +7902,35 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_94 = lean_ctor_get(x_91, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_91, 1); -lean_inc(x_95); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_96 = x_91; +x_155 = lean_ctor_get(x_127, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_127, 1); +lean_inc(x_156); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_157 = x_127; } else { - lean_dec_ref(x_91); - x_96 = lean_box(0); + lean_dec_ref(x_127); + x_157 = lean_box(0); } -if (lean_is_scalar(x_96)) { - x_97 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_157)) { + x_158 = lean_alloc_ctor(1, 2, 0); } else { - x_97 = x_96; -} -lean_ctor_set(x_97, 0, x_94); -lean_ctor_set(x_97, 1, x_95); -return x_97; + x_158 = x_157; } +lean_ctor_set(x_158, 0, x_155); +lean_ctor_set(x_158, 1, x_156); +return x_158; } -else -{ -lean_object* x_98; lean_object* x_99; -x_98 = lean_ctor_get(x_1, 5); -lean_inc(x_98); -x_99 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_98, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_87); -if (lean_obj_tag(x_99) == 0) -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_100 = lean_ctor_get(x_99, 0); -lean_inc(x_100); -x_101 = lean_ctor_get(x_99, 1); -lean_inc(x_101); -lean_dec(x_99); -x_102 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__5; -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_100); -x_104 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__7; -x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_MessageData_ofExpr(x_39); -x_107 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -x_108 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9; -x_109 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_indentExpr(x_36); -x_111 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -x_112 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_113 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -x_114 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_42, x_113, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_101); -x_115 = lean_ctor_get(x_114, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 1); -lean_inc(x_116); -lean_dec(x_114); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -x_117 = lean_apply_12(x_88, x_115, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_116); -if (lean_obj_tag(x_117) == 0) -{ -lean_object* x_118; lean_object* x_119; -x_118 = lean_ctor_get(x_117, 0); -lean_inc(x_118); -x_119 = lean_ctor_get(x_117, 1); -lean_inc(x_119); -lean_dec(x_117); -x_27 = x_118; -x_28 = x_119; -goto block_35; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +lean_dec(x_45); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_36); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -5099,31 +7944,33 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_120 = lean_ctor_get(x_117, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_117, 1); -lean_inc(x_121); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_122 = x_117; +x_159 = lean_ctor_get(x_124, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_124, 1); +lean_inc(x_160); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_161 = x_124; } else { - lean_dec_ref(x_117); - x_122 = lean_box(0); + lean_dec_ref(x_124); + x_161 = lean_box(0); } -if (lean_is_scalar(x_122)) { - x_123 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_161)) { + x_162 = lean_alloc_ctor(1, 2, 0); } else { - x_123 = x_122; + x_162 = x_161; } -lean_ctor_set(x_123, 0, x_120); -lean_ctor_set(x_123, 1, x_121); -return x_123; +lean_ctor_set(x_162, 0, x_159); +lean_ctor_set(x_162, 1, x_160); +return x_162; } } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +lean_dec(x_45); +lean_dec(x_42); lean_dec(x_39); lean_dec(x_36); lean_dec(x_22); @@ -5139,52 +7986,133 @@ lean_dec(x_13); lean_dec(x_10); lean_dec(x_7); lean_dec(x_1); -x_124 = lean_ctor_get(x_99, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_99, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_99)) { - lean_ctor_release(x_99, 0); - lean_ctor_release(x_99, 1); - x_126 = x_99; +x_163 = lean_ctor_get(x_121, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_121, 1); +lean_inc(x_164); +if (lean_is_exclusive(x_121)) { + lean_ctor_release(x_121, 0); + lean_ctor_release(x_121, 1); + x_165 = x_121; } else { - lean_dec_ref(x_99); - x_126 = lean_box(0); + lean_dec_ref(x_121); + x_165 = lean_box(0); } -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_165)) { + x_166 = lean_alloc_ctor(1, 2, 0); } else { - x_127 = x_126; + x_166 = x_165; } -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_ctor_set(x_166, 0, x_163); +lean_ctor_set(x_166, 1, x_164); +return x_166; } } } } else { -lean_object* x_128; +lean_object* x_167; +lean_dec(x_45); +lean_dec(x_42); lean_dec(x_39); lean_dec(x_36); lean_inc(x_7); -x_128 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_128, 0, x_7); -x_27 = x_128; -x_28 = x_41; +x_167 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_167, 0, x_7); +x_27 = x_167; +x_28 = x_48; goto block_35; } } } else { -lean_object* x_163; +uint8_t x_187; +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_187 = !lean_is_exclusive(x_44); +if (x_187 == 0) +{ +return x_44; +} +else +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_188 = lean_ctor_get(x_44, 0); +x_189 = lean_ctor_get(x_44, 1); +lean_inc(x_189); +lean_inc(x_188); +lean_dec(x_44); +x_190 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +return x_190; +} +} +} +else +{ +uint8_t x_191; +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_1); +x_191 = !lean_is_exclusive(x_41); +if (x_191 == 0) +{ +return x_41; +} +else +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_ctor_get(x_41, 0); +x_193 = lean_ctor_get(x_41, 1); +lean_inc(x_193); +lean_inc(x_192); +lean_dec(x_41); +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +return x_194; +} +} +} +else +{ +lean_object* x_195; lean_dec(x_36); lean_inc(x_7); -x_163 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_163, 0, x_7); -x_27 = x_163; +x_195 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_195, 0, x_7); +x_27 = x_195; x_28 = x_23; goto block_35; } @@ -5438,56 +8366,63 @@ uint8_t x_72; x_72 = !lean_is_exclusive(x_64); if (x_72 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_object* x_73; lean_object* x_74; lean_object* x_75; x_73 = lean_ctor_get(x_64, 1); x_74 = lean_ctor_get(x_64, 0); lean_dec(x_74); -x_75 = lean_ctor_get(x_1, 5); -lean_inc(x_75); -x_76 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_73); -if (lean_obj_tag(x_76) == 0) +x_75 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_73); +if (lean_obj_tag(x_75) == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_77 = lean_ctor_get(x_76, 0); +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_ctor_get(x_1, 5); lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +x_78 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_77, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_76); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_78, 1); +lean_inc(x_80); +lean_dec(x_78); +x_81 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; lean_ctor_set_tag(x_64, 7); -lean_ctor_set(x_64, 1, x_77); -lean_ctor_set(x_64, 0, x_79); -x_80 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set(x_64, 1, x_79); +lean_ctor_set(x_64, 0, x_81); +x_82 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; lean_ctor_set_tag(x_49, 7); -lean_ctor_set(x_49, 1, x_80); +lean_ctor_set(x_49, 1, x_82); lean_ctor_set(x_49, 0, x_64); -x_81 = l_Lean_indentExpr(x_57); -x_82 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_82, 0, x_49); -lean_ctor_set(x_82, 1, x_81); -x_83 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_80); -x_84 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_83, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_78); -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_84, 1); -lean_inc(x_86); -lean_dec(x_84); -x_87 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_85, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_86); -lean_dec(x_85); -x_88 = lean_ctor_get(x_87, 0); +x_83 = l_Lean_indentExpr(x_57); +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_49); +lean_ctor_set(x_84, 1, x_83); +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +x_86 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_85, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_80); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); +lean_dec(x_86); +x_89 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_87, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_88); lean_dec(x_87); -x_23 = x_88; -x_24 = x_89; +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); +x_23 = x_90; +x_24 = x_91; goto block_31; } else { -uint8_t x_90; +uint8_t x_92; lean_free_object(x_64); lean_dec(x_57); lean_free_object(x_49); @@ -5498,78 +8433,118 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_90 = !lean_is_exclusive(x_76); -if (x_90 == 0) +x_92 = !lean_is_exclusive(x_78); +if (x_92 == 0) { -return x_76; +return x_78; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_76, 0); -x_92 = lean_ctor_get(x_76, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_76); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_78, 0); +x_94 = lean_ctor_get(x_78, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_78); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_64, 1); -lean_inc(x_94); -lean_dec(x_64); -x_95 = lean_ctor_get(x_1, 5); -lean_inc(x_95); -x_96 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_95, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_94); -if (lean_obj_tag(x_96) == 0) +uint8_t x_96; +lean_free_object(x_64); +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_96 = !lean_is_exclusive(x_75); +if (x_96 == 0) { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); +return x_75; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_75, 0); +x_98 = lean_ctor_get(x_75, 1); lean_inc(x_98); -lean_dec(x_96); -x_99 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -x_100 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -x_101 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_49, 7); -lean_ctor_set(x_49, 1, x_101); -lean_ctor_set(x_49, 0, x_100); -x_102 = l_Lean_indentExpr(x_57); -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_49); -lean_ctor_set(x_103, 1, x_102); -x_104 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_101); -x_105 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_104, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_98); -x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_97); +lean_dec(x_75); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; +} +} +} +else +{ +lean_object* x_100; lean_object* x_101; +x_100 = lean_ctor_get(x_64, 1); +lean_inc(x_100); +lean_dec(x_64); +x_101 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_100); +if (lean_obj_tag(x_101) == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +lean_dec(x_101); +x_103 = lean_ctor_get(x_1, 5); +lean_inc(x_103); +x_104 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_103, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_102); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); lean_inc(x_106); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -lean_dec(x_105); -x_108 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_106, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_107); -lean_dec(x_106); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -lean_dec(x_108); -x_23 = x_109; -x_24 = x_110; +lean_dec(x_104); +x_107 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +x_108 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_105); +x_109 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_49, 7); +lean_ctor_set(x_49, 1, x_109); +lean_ctor_set(x_49, 0, x_108); +x_110 = l_Lean_indentExpr(x_57); +x_111 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_111, 0, x_49); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_109); +x_113 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_63, x_112, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_106); +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_113, 1); +lean_inc(x_115); +lean_dec(x_113); +x_116 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_114, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_115); +lean_dec(x_114); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); +lean_dec(x_116); +x_23 = x_117; +x_24 = x_118; goto block_31; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_dec(x_57); lean_free_object(x_49); lean_dec(x_33); @@ -5579,50 +8554,84 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_111 = lean_ctor_get(x_96, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_96, 1); -lean_inc(x_112); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_113 = x_96; +x_119 = lean_ctor_get(x_104, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_104, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_121 = x_104; +} else { + lean_dec_ref(x_104); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 2, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_119); +lean_ctor_set(x_122, 1, x_120); +return x_122; +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_dec(x_57); +lean_free_object(x_49); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_123 = lean_ctor_get(x_101, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_101, 1); +lean_inc(x_124); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_125 = x_101; } else { - lean_dec_ref(x_96); - x_113 = lean_box(0); + lean_dec_ref(x_101); + x_125 = lean_box(0); } -if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_125)) { + x_126 = lean_alloc_ctor(1, 2, 0); } else { - x_114 = x_113; + x_126 = x_125; } -lean_ctor_set(x_114, 0, x_111); -lean_ctor_set(x_114, 1, x_112); -return x_114; +lean_ctor_set(x_126, 0, x_123); +lean_ctor_set(x_126, 1, x_124); +return x_126; } } } } else { -lean_object* x_115; lean_object* x_116; +lean_object* x_127; lean_object* x_128; lean_dec(x_57); lean_free_object(x_49); -x_115 = lean_ctor_get(x_59, 1); -lean_inc(x_115); +x_127 = lean_ctor_get(x_59, 1); +lean_inc(x_127); lean_dec(x_59); lean_inc(x_3); lean_ctor_set(x_8, 0, x_3); -x_116 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_116, 0, x_8); -x_23 = x_116; -x_24 = x_115; +x_128 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_128, 0, x_8); +x_23 = x_128; +x_24 = x_127; goto block_31; } } else { -uint8_t x_117; +uint8_t x_129; lean_dec(x_57); lean_free_object(x_49); lean_dec(x_33); @@ -5633,29 +8642,29 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_117 = !lean_is_exclusive(x_59); -if (x_117 == 0) +x_129 = !lean_is_exclusive(x_59); +if (x_129 == 0) { return x_59; } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_59, 0); -x_119 = lean_ctor_get(x_59, 1); -lean_inc(x_119); -lean_inc(x_118); +lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_59, 0); +x_131 = lean_ctor_get(x_59, 1); +lean_inc(x_131); +lean_inc(x_130); lean_dec(x_59); -x_120 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set(x_120, 1, x_119); -return x_120; +x_132 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +return x_132; } } } else { -uint8_t x_121; +uint8_t x_133; lean_free_object(x_49); lean_dec(x_33); lean_free_object(x_8); @@ -5666,36 +8675,36 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_121 = !lean_is_exclusive(x_56); -if (x_121 == 0) +x_133 = !lean_is_exclusive(x_56); +if (x_133 == 0) { return x_56; } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_56, 0); -x_123 = lean_ctor_get(x_56, 1); -lean_inc(x_123); -lean_inc(x_122); +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_56, 0); +x_135 = lean_ctor_get(x_56, 1); +lean_inc(x_135); +lean_inc(x_134); lean_dec(x_56); -x_124 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); -return x_124; +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +return x_136; } } } else { -lean_object* x_125; +lean_object* x_137; lean_free_object(x_49); lean_dec(x_22); lean_inc(x_3); lean_ctor_set(x_8, 0, x_3); -x_125 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_125, 0, x_8); -x_23 = x_125; +x_137 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_137, 0, x_8); +x_23 = x_137; x_24 = x_52; goto block_31; } @@ -5703,158 +8712,165 @@ goto block_31; } else { -lean_object* x_126; lean_object* x_127; uint8_t x_128; -x_126 = lean_ctor_get(x_49, 0); -x_127 = lean_ctor_get(x_49, 1); -lean_inc(x_127); -lean_inc(x_126); +lean_object* x_138; lean_object* x_139; uint8_t x_140; +x_138 = lean_ctor_get(x_49, 0); +x_139 = lean_ctor_get(x_49, 1); +lean_inc(x_139); +lean_inc(x_138); lean_dec(x_49); -x_128 = l_Lean_BinderInfo_isInstImplicit(x_45); -if (x_128 == 0) +x_140 = l_Lean_BinderInfo_isInstImplicit(x_45); +if (x_140 == 0) { -lean_object* x_129; -lean_dec(x_126); +lean_object* x_141; +lean_dec(x_138); lean_dec(x_22); lean_inc(x_3); lean_ctor_set(x_8, 0, x_3); -x_129 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_129, 0, x_8); -x_23 = x_129; -x_24 = x_127; +x_141 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_141, 0, x_8); +x_23 = x_141; +x_24 = x_139; goto block_31; } else { -uint8_t x_130; -x_130 = lean_unbox(x_126); -lean_dec(x_126); -if (x_130 == 0) +uint8_t x_142; +x_142 = lean_unbox(x_138); +lean_dec(x_138); +if (x_142 == 0) { -lean_object* x_131; +lean_object* x_143; lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_22); -x_131 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_127); -if (lean_obj_tag(x_131) == 0) +x_143 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_139); +if (lean_obj_tag(x_143) == 0) { -lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_132); -x_134 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_132, x_15, x_16, x_17, x_18, x_133); -if (lean_obj_tag(x_134) == 0) +lean_inc(x_144); +x_146 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_144, x_15, x_16, x_17, x_18, x_145); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_135; uint8_t x_136; -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_unbox(x_135); -lean_dec(x_135); -if (x_136 == 0) +lean_object* x_147; uint8_t x_148; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_unbox(x_147); +lean_dec(x_147); +if (x_148 == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_free_object(x_8); -x_137 = lean_ctor_get(x_134, 1); -lean_inc(x_137); -lean_dec(x_134); -x_138 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_139 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_138, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_137); -x_140 = lean_ctor_get(x_139, 0); -lean_inc(x_140); -x_141 = lean_unbox(x_140); -lean_dec(x_140); -if (x_141 == 0) +x_149 = lean_ctor_get(x_146, 1); +lean_inc(x_149); +lean_dec(x_146); +x_150 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_151 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_150, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_149); +x_152 = lean_ctor_get(x_151, 0); +lean_inc(x_152); +x_153 = lean_unbox(x_152); +lean_dec(x_152); +if (x_153 == 0) { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_132); -x_142 = lean_ctor_get(x_139, 1); -lean_inc(x_142); -lean_dec(x_139); -x_143 = lean_box(0); -x_144 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_143, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_142); -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_144, 1); -lean_inc(x_146); +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_dec(x_144); -x_23 = x_145; -x_24 = x_146; +x_154 = lean_ctor_get(x_151, 1); +lean_inc(x_154); +lean_dec(x_151); +x_155 = lean_box(0); +x_156 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_155, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_154); +x_157 = lean_ctor_get(x_156, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_156, 1); +lean_inc(x_158); +lean_dec(x_156); +x_23 = x_157; +x_24 = x_158; goto block_31; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_147 = lean_ctor_get(x_139, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_148 = x_139; +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_151, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_151)) { + lean_ctor_release(x_151, 0); + lean_ctor_release(x_151, 1); + x_160 = x_151; } else { - lean_dec_ref(x_139); - x_148 = lean_box(0); + lean_dec_ref(x_151); + x_160 = lean_box(0); } -x_149 = lean_ctor_get(x_1, 5); -lean_inc(x_149); -x_150 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_149, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_147); -if (lean_obj_tag(x_150) == 0) +x_161 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_159); +if (lean_obj_tag(x_161) == 0) { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec(x_150); -x_153 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -if (lean_is_scalar(x_148)) { - x_154 = lean_alloc_ctor(7, 2, 0); -} else { - x_154 = x_148; - lean_ctor_set_tag(x_154, 7); -} -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_151); -x_155 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_156 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_156, 0, x_154); -lean_ctor_set(x_156, 1, x_155); -x_157 = l_Lean_indentExpr(x_132); -x_158 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -x_159 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_159, 0, x_158); -lean_ctor_set(x_159, 1, x_155); -x_160 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_138, x_159, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_152); -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_160, 1); +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_161, 1); lean_inc(x_162); -lean_dec(x_160); -x_163 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_161, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_162); lean_dec(x_161); -x_164 = lean_ctor_get(x_163, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_163, 1); +x_163 = lean_ctor_get(x_1, 5); +lean_inc(x_163); +x_164 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_163, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_162); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_165 = lean_ctor_get(x_164, 0); lean_inc(x_165); -lean_dec(x_163); -x_23 = x_164; -x_24 = x_165; +x_166 = lean_ctor_get(x_164, 1); +lean_inc(x_166); +lean_dec(x_164); +x_167 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_160)) { + x_168 = lean_alloc_ctor(7, 2, 0); +} else { + x_168 = x_160; + lean_ctor_set_tag(x_168, 7); +} +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_165); +x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_170 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_169); +x_171 = l_Lean_indentExpr(x_144); +x_172 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +x_173 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set(x_173, 1, x_169); +x_174 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_150, x_173, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_166); +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_174, 1); +lean_inc(x_176); +lean_dec(x_174); +x_177 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_33, x_175, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_176); +lean_dec(x_175); +x_178 = lean_ctor_get(x_177, 0); +lean_inc(x_178); +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +lean_dec(x_177); +x_23 = x_178; +x_24 = x_179; goto block_31; } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; -lean_dec(x_148); -lean_dec(x_132); +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_dec(x_160); +lean_dec(x_144); lean_dec(x_33); lean_dec(x_18); lean_dec(x_17); @@ -5862,49 +8878,83 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_166 = lean_ctor_get(x_150, 0); -lean_inc(x_166); -x_167 = lean_ctor_get(x_150, 1); -lean_inc(x_167); -if (lean_is_exclusive(x_150)) { - lean_ctor_release(x_150, 0); - lean_ctor_release(x_150, 1); - x_168 = x_150; +x_180 = lean_ctor_get(x_164, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_164, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_182 = x_164; } else { - lean_dec_ref(x_150); - x_168 = lean_box(0); + lean_dec_ref(x_164); + x_182 = lean_box(0); } -if (lean_is_scalar(x_168)) { - x_169 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_182)) { + x_183 = lean_alloc_ctor(1, 2, 0); } else { - x_169 = x_168; + x_183 = x_182; } -lean_ctor_set(x_169, 0, x_166); -lean_ctor_set(x_169, 1, x_167); -return x_169; +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_181); +return x_183; +} +} +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +lean_dec(x_160); +lean_dec(x_144); +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_184 = lean_ctor_get(x_161, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_161, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_186 = x_161; +} else { + lean_dec_ref(x_161); + x_186 = lean_box(0); +} +if (lean_is_scalar(x_186)) { + x_187 = lean_alloc_ctor(1, 2, 0); +} else { + x_187 = x_186; +} +lean_ctor_set(x_187, 0, x_184); +lean_ctor_set(x_187, 1, x_185); +return x_187; } } } else { -lean_object* x_170; lean_object* x_171; -lean_dec(x_132); -x_170 = lean_ctor_get(x_134, 1); -lean_inc(x_170); -lean_dec(x_134); +lean_object* x_188; lean_object* x_189; +lean_dec(x_144); +x_188 = lean_ctor_get(x_146, 1); +lean_inc(x_188); +lean_dec(x_146); lean_inc(x_3); lean_ctor_set(x_8, 0, x_3); -x_171 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_171, 0, x_8); -x_23 = x_171; -x_24 = x_170; +x_189 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_189, 0, x_8); +x_23 = x_189; +x_24 = x_188; goto block_31; } } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -lean_dec(x_132); +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_144); lean_dec(x_33); lean_free_object(x_8); lean_dec(x_18); @@ -5913,31 +8963,31 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_172 = lean_ctor_get(x_134, 0); -lean_inc(x_172); -x_173 = lean_ctor_get(x_134, 1); -lean_inc(x_173); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_174 = x_134; +x_190 = lean_ctor_get(x_146, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_146, 1); +lean_inc(x_191); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_192 = x_146; } else { - lean_dec_ref(x_134); - x_174 = lean_box(0); + lean_dec_ref(x_146); + x_192 = lean_box(0); } -if (lean_is_scalar(x_174)) { - x_175 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_192)) { + x_193 = lean_alloc_ctor(1, 2, 0); } else { - x_175 = x_174; + x_193 = x_192; } -lean_ctor_set(x_175, 0, x_172); -lean_ctor_set(x_175, 1, x_173); -return x_175; +lean_ctor_set(x_193, 0, x_190); +lean_ctor_set(x_193, 1, x_191); +return x_193; } } else { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_dec(x_33); lean_free_object(x_8); lean_dec(x_22); @@ -5947,38 +8997,38 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_176 = lean_ctor_get(x_131, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_131, 1); -lean_inc(x_177); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_178 = x_131; +x_194 = lean_ctor_get(x_143, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_143, 1); +lean_inc(x_195); +if (lean_is_exclusive(x_143)) { + lean_ctor_release(x_143, 0); + lean_ctor_release(x_143, 1); + x_196 = x_143; } else { - lean_dec_ref(x_131); - x_178 = lean_box(0); + lean_dec_ref(x_143); + x_196 = lean_box(0); } -if (lean_is_scalar(x_178)) { - x_179 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_196)) { + x_197 = lean_alloc_ctor(1, 2, 0); } else { - x_179 = x_178; + x_197 = x_196; } -lean_ctor_set(x_179, 0, x_176); -lean_ctor_set(x_179, 1, x_177); -return x_179; +lean_ctor_set(x_197, 0, x_194); +lean_ctor_set(x_197, 1, x_195); +return x_197; } } else { -lean_object* x_180; +lean_object* x_198; lean_dec(x_22); lean_inc(x_3); lean_ctor_set(x_8, 0, x_3); -x_180 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_180, 0, x_8); -x_23 = x_180; -x_24 = x_127; +x_198 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_198, 0, x_8); +x_23 = x_198; +x_24 = x_139; goto block_31; } } @@ -5986,242 +9036,284 @@ goto block_31; } else { -lean_object* x_181; uint8_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; +lean_object* x_199; uint8_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; lean_dec(x_33); -x_181 = lean_array_fget(x_35, x_36); -x_182 = lean_unbox(x_181); -lean_dec(x_181); -x_183 = lean_unsigned_to_nat(1u); -x_184 = lean_nat_add(x_36, x_183); +x_199 = lean_array_fget(x_35, x_36); +x_200 = lean_unbox(x_199); +lean_dec(x_199); +x_201 = lean_unsigned_to_nat(1u); +x_202 = lean_nat_add(x_36, x_201); lean_dec(x_36); -x_185 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_185, 0, x_35); -lean_ctor_set(x_185, 1, x_184); -lean_ctor_set(x_185, 2, x_37); -x_186 = l_Lean_Expr_mvarId_x21(x_22); -x_187 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_186, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_186); -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_187, 1); -lean_inc(x_189); -if (lean_is_exclusive(x_187)) { - lean_ctor_release(x_187, 0); - lean_ctor_release(x_187, 1); - x_190 = x_187; +x_203 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_203, 0, x_35); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_37); +x_204 = l_Lean_Expr_mvarId_x21(x_22); +x_205 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_204, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_204); +x_206 = lean_ctor_get(x_205, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_205, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_205)) { + lean_ctor_release(x_205, 0); + lean_ctor_release(x_205, 1); + x_208 = x_205; } else { - lean_dec_ref(x_187); - x_190 = lean_box(0); + lean_dec_ref(x_205); + x_208 = lean_box(0); } -x_191 = l_Lean_BinderInfo_isInstImplicit(x_182); -if (x_191 == 0) +x_209 = l_Lean_BinderInfo_isInstImplicit(x_200); +if (x_209 == 0) { -lean_object* x_192; -lean_dec(x_190); -lean_dec(x_188); +lean_object* x_210; +lean_dec(x_208); +lean_dec(x_206); lean_dec(x_22); lean_inc(x_3); -lean_ctor_set(x_8, 1, x_185); +lean_ctor_set(x_8, 1, x_203); lean_ctor_set(x_8, 0, x_3); -x_192 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_192, 0, x_8); -x_23 = x_192; -x_24 = x_189; +x_210 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_210, 0, x_8); +x_23 = x_210; +x_24 = x_207; goto block_31; } else { -uint8_t x_193; -x_193 = lean_unbox(x_188); -lean_dec(x_188); -if (x_193 == 0) +uint8_t x_211; +x_211 = lean_unbox(x_206); +lean_dec(x_206); +if (x_211 == 0) { -lean_object* x_194; +lean_object* x_212; lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_22); -x_194 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_189); -if (lean_obj_tag(x_194) == 0) +x_212 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_207); +if (lean_obj_tag(x_212) == 0) { -lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_195 = lean_ctor_get(x_194, 0); -lean_inc(x_195); -x_196 = lean_ctor_get(x_194, 1); -lean_inc(x_196); -lean_dec(x_194); +lean_object* x_213; lean_object* x_214; lean_object* x_215; +x_213 = lean_ctor_get(x_212, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_212, 1); +lean_inc(x_214); +lean_dec(x_212); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_195); -x_197 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_195, x_15, x_16, x_17, x_18, x_196); -if (lean_obj_tag(x_197) == 0) +lean_inc(x_213); +x_215 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_213, x_15, x_16, x_17, x_18, x_214); +if (lean_obj_tag(x_215) == 0) { -lean_object* x_198; uint8_t x_199; -x_198 = lean_ctor_get(x_197, 0); -lean_inc(x_198); -x_199 = lean_unbox(x_198); -lean_dec(x_198); -if (x_199 == 0) +lean_object* x_216; uint8_t x_217; +x_216 = lean_ctor_get(x_215, 0); +lean_inc(x_216); +x_217 = lean_unbox(x_216); +lean_dec(x_216); +if (x_217 == 0) { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; uint8_t x_222; lean_free_object(x_8); -x_200 = lean_ctor_get(x_197, 1); -lean_inc(x_200); -lean_dec(x_197); -x_201 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_202 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_201, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_200); -x_203 = lean_ctor_get(x_202, 0); -lean_inc(x_203); -x_204 = lean_unbox(x_203); -lean_dec(x_203); -if (x_204 == 0) +x_218 = lean_ctor_get(x_215, 1); +lean_inc(x_218); +lean_dec(x_215); +x_219 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_220 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_219, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_218); +x_221 = lean_ctor_get(x_220, 0); +lean_inc(x_221); +x_222 = lean_unbox(x_221); +lean_dec(x_221); +if (x_222 == 0) { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_195); -lean_dec(x_190); -x_205 = lean_ctor_get(x_202, 1); -lean_inc(x_205); -lean_dec(x_202); -x_206 = lean_box(0); -x_207 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_185, x_206, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_205); -x_208 = lean_ctor_get(x_207, 0); -lean_inc(x_208); -x_209 = lean_ctor_get(x_207, 1); -lean_inc(x_209); -lean_dec(x_207); -x_23 = x_208; -x_24 = x_209; +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; +lean_dec(x_213); +lean_dec(x_208); +x_223 = lean_ctor_get(x_220, 1); +lean_inc(x_223); +lean_dec(x_220); +x_224 = lean_box(0); +x_225 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_203, x_224, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_223); +x_226 = lean_ctor_get(x_225, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_225, 1); +lean_inc(x_227); +lean_dec(x_225); +x_23 = x_226; +x_24 = x_227; goto block_31; } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_210 = lean_ctor_get(x_202, 1); -lean_inc(x_210); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - lean_ctor_release(x_202, 1); - x_211 = x_202; +lean_object* x_228; lean_object* x_229; lean_object* x_230; +x_228 = lean_ctor_get(x_220, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_220)) { + lean_ctor_release(x_220, 0); + lean_ctor_release(x_220, 1); + x_229 = x_220; +} else { + lean_dec_ref(x_220); + x_229 = lean_box(0); +} +x_230 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_228); +if (lean_obj_tag(x_230) == 0) +{ +lean_object* x_231; lean_object* x_232; lean_object* x_233; +x_231 = lean_ctor_get(x_230, 1); +lean_inc(x_231); +lean_dec(x_230); +x_232 = lean_ctor_get(x_1, 5); +lean_inc(x_232); +x_233 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_232, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_231); +if (lean_obj_tag(x_233) == 0) +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_233, 1); +lean_inc(x_235); +lean_dec(x_233); +x_236 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_229)) { + x_237 = lean_alloc_ctor(7, 2, 0); +} else { + x_237 = x_229; + lean_ctor_set_tag(x_237, 7); +} +lean_ctor_set(x_237, 0, x_236); +lean_ctor_set(x_237, 1, x_234); +x_238 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +if (lean_is_scalar(x_208)) { + x_239 = lean_alloc_ctor(7, 2, 0); } else { - lean_dec_ref(x_202); - x_211 = lean_box(0); + x_239 = x_208; + lean_ctor_set_tag(x_239, 7); +} +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +x_240 = l_Lean_indentExpr(x_213); +x_241 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_241, 0, x_239); +lean_ctor_set(x_241, 1, x_240); +x_242 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_242, 0, x_241); +lean_ctor_set(x_242, 1, x_238); +x_243 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_219, x_242, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_235); +x_244 = lean_ctor_get(x_243, 0); +lean_inc(x_244); +x_245 = lean_ctor_get(x_243, 1); +lean_inc(x_245); +lean_dec(x_243); +x_246 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_203, x_244, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_245); +lean_dec(x_244); +x_247 = lean_ctor_get(x_246, 0); +lean_inc(x_247); +x_248 = lean_ctor_get(x_246, 1); +lean_inc(x_248); +lean_dec(x_246); +x_23 = x_247; +x_24 = x_248; +goto block_31; } -x_212 = lean_ctor_get(x_1, 5); -lean_inc(x_212); -x_213 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_212, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_210); -if (lean_obj_tag(x_213) == 0) +else { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); +lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +lean_dec(x_229); lean_dec(x_213); -x_216 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -if (lean_is_scalar(x_211)) { - x_217 = lean_alloc_ctor(7, 2, 0); -} else { - x_217 = x_211; - lean_ctor_set_tag(x_217, 7); -} -lean_ctor_set(x_217, 0, x_216); -lean_ctor_set(x_217, 1, x_214); -x_218 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -if (lean_is_scalar(x_190)) { - x_219 = lean_alloc_ctor(7, 2, 0); -} else { - x_219 = x_190; - lean_ctor_set_tag(x_219, 7); -} -lean_ctor_set(x_219, 0, x_217); -lean_ctor_set(x_219, 1, x_218); -x_220 = l_Lean_indentExpr(x_195); -x_221 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -x_222 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_222, 0, x_221); -lean_ctor_set(x_222, 1, x_218); -x_223 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_201, x_222, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_215); -x_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_223, 1); -lean_inc(x_225); -lean_dec(x_223); -x_226 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_185, x_224, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_225); -lean_dec(x_224); -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_226, 1); -lean_inc(x_228); -lean_dec(x_226); -x_23 = x_227; -x_24 = x_228; -goto block_31; +lean_dec(x_208); +lean_dec(x_203); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_249 = lean_ctor_get(x_233, 0); +lean_inc(x_249); +x_250 = lean_ctor_get(x_233, 1); +lean_inc(x_250); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_251 = x_233; +} else { + lean_dec_ref(x_233); + x_251 = lean_box(0); +} +if (lean_is_scalar(x_251)) { + x_252 = lean_alloc_ctor(1, 2, 0); +} else { + x_252 = x_251; +} +lean_ctor_set(x_252, 0, x_249); +lean_ctor_set(x_252, 1, x_250); +return x_252; +} } else { -lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -lean_dec(x_211); -lean_dec(x_195); -lean_dec(x_190); -lean_dec(x_185); +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +lean_dec(x_229); +lean_dec(x_213); +lean_dec(x_208); +lean_dec(x_203); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_229 = lean_ctor_get(x_213, 0); -lean_inc(x_229); -x_230 = lean_ctor_get(x_213, 1); -lean_inc(x_230); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_231 = x_213; +x_253 = lean_ctor_get(x_230, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_230, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_230)) { + lean_ctor_release(x_230, 0); + lean_ctor_release(x_230, 1); + x_255 = x_230; } else { - lean_dec_ref(x_213); - x_231 = lean_box(0); + lean_dec_ref(x_230); + x_255 = lean_box(0); } -if (lean_is_scalar(x_231)) { - x_232 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(1, 2, 0); } else { - x_232 = x_231; + x_256 = x_255; } -lean_ctor_set(x_232, 0, x_229); -lean_ctor_set(x_232, 1, x_230); -return x_232; +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_254); +return x_256; } } } else { -lean_object* x_233; lean_object* x_234; -lean_dec(x_195); -lean_dec(x_190); -x_233 = lean_ctor_get(x_197, 1); -lean_inc(x_233); -lean_dec(x_197); +lean_object* x_257; lean_object* x_258; +lean_dec(x_213); +lean_dec(x_208); +x_257 = lean_ctor_get(x_215, 1); +lean_inc(x_257); +lean_dec(x_215); lean_inc(x_3); -lean_ctor_set(x_8, 1, x_185); +lean_ctor_set(x_8, 1, x_203); lean_ctor_set(x_8, 0, x_3); -x_234 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_234, 0, x_8); -x_23 = x_234; -x_24 = x_233; +x_258 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_258, 0, x_8); +x_23 = x_258; +x_24 = x_257; goto block_31; } } else { -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_dec(x_195); -lean_dec(x_190); -lean_dec(x_185); +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +lean_dec(x_213); +lean_dec(x_208); +lean_dec(x_203); lean_free_object(x_8); lean_dec(x_18); lean_dec(x_17); @@ -6229,33 +9321,33 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_235 = lean_ctor_get(x_197, 0); -lean_inc(x_235); -x_236 = lean_ctor_get(x_197, 1); -lean_inc(x_236); -if (lean_is_exclusive(x_197)) { - lean_ctor_release(x_197, 0); - lean_ctor_release(x_197, 1); - x_237 = x_197; +x_259 = lean_ctor_get(x_215, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_215, 1); +lean_inc(x_260); +if (lean_is_exclusive(x_215)) { + lean_ctor_release(x_215, 0); + lean_ctor_release(x_215, 1); + x_261 = x_215; } else { - lean_dec_ref(x_197); - x_237 = lean_box(0); + lean_dec_ref(x_215); + x_261 = lean_box(0); } -if (lean_is_scalar(x_237)) { - x_238 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_261)) { + x_262 = lean_alloc_ctor(1, 2, 0); } else { - x_238 = x_237; + x_262 = x_261; } -lean_ctor_set(x_238, 0, x_235); -lean_ctor_set(x_238, 1, x_236); -return x_238; +lean_ctor_set(x_262, 0, x_259); +lean_ctor_set(x_262, 1, x_260); +return x_262; } } else { -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; -lean_dec(x_190); -lean_dec(x_185); +lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +lean_dec(x_208); +lean_dec(x_203); lean_free_object(x_8); lean_dec(x_22); lean_dec(x_18); @@ -6264,40 +9356,40 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_239 = lean_ctor_get(x_194, 0); -lean_inc(x_239); -x_240 = lean_ctor_get(x_194, 1); -lean_inc(x_240); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - x_241 = x_194; +x_263 = lean_ctor_get(x_212, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_212, 1); +lean_inc(x_264); +if (lean_is_exclusive(x_212)) { + lean_ctor_release(x_212, 0); + lean_ctor_release(x_212, 1); + x_265 = x_212; } else { - lean_dec_ref(x_194); - x_241 = lean_box(0); + lean_dec_ref(x_212); + x_265 = lean_box(0); } -if (lean_is_scalar(x_241)) { - x_242 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_265)) { + x_266 = lean_alloc_ctor(1, 2, 0); } else { - x_242 = x_241; + x_266 = x_265; } -lean_ctor_set(x_242, 0, x_239); -lean_ctor_set(x_242, 1, x_240); -return x_242; +lean_ctor_set(x_266, 0, x_263); +lean_ctor_set(x_266, 1, x_264); +return x_266; } } else { -lean_object* x_243; -lean_dec(x_190); +lean_object* x_267; +lean_dec(x_208); lean_dec(x_22); lean_inc(x_3); -lean_ctor_set(x_8, 1, x_185); +lean_ctor_set(x_8, 1, x_203); lean_ctor_set(x_8, 0, x_3); -x_243 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_243, 0, x_8); -x_23 = x_243; -x_24 = x_189; +x_267 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_267, 0, x_8); +x_23 = x_267; +x_24 = x_207; goto block_31; } } @@ -6306,318 +9398,360 @@ goto block_31; } else { -lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; -x_244 = lean_ctor_get(x_8, 1); -lean_inc(x_244); +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; uint8_t x_272; +x_268 = lean_ctor_get(x_8, 1); +lean_inc(x_268); lean_dec(x_8); -x_245 = lean_ctor_get(x_244, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_244, 1); -lean_inc(x_246); -x_247 = lean_ctor_get(x_244, 2); -lean_inc(x_247); -x_248 = lean_nat_dec_lt(x_246, x_247); -if (x_248 == 0) +x_269 = lean_ctor_get(x_268, 0); +lean_inc(x_269); +x_270 = lean_ctor_get(x_268, 1); +lean_inc(x_270); +x_271 = lean_ctor_get(x_268, 2); +lean_inc(x_271); +x_272 = lean_nat_dec_lt(x_270, x_271); +if (x_272 == 0) { -lean_object* x_249; lean_object* x_250; -lean_dec(x_247); -lean_dec(x_246); -lean_dec(x_245); +lean_object* x_273; lean_object* x_274; +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); lean_dec(x_22); lean_inc(x_3); -x_249 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_249, 0, x_3); -lean_ctor_set(x_249, 1, x_244); -x_250 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_250, 0, x_249); -x_23 = x_250; +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_3); +lean_ctor_set(x_273, 1, x_268); +x_274 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_274, 0, x_273); +x_23 = x_274; x_24 = x_19; goto block_31; } else { -lean_object* x_251; lean_object* x_252; uint8_t x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; uint8_t x_262; -if (lean_is_exclusive(x_244)) { - lean_ctor_release(x_244, 0); - lean_ctor_release(x_244, 1); - lean_ctor_release(x_244, 2); - x_251 = x_244; +lean_object* x_275; lean_object* x_276; uint8_t x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; uint8_t x_286; +if (lean_is_exclusive(x_268)) { + lean_ctor_release(x_268, 0); + lean_ctor_release(x_268, 1); + lean_ctor_release(x_268, 2); + x_275 = x_268; } else { - lean_dec_ref(x_244); - x_251 = lean_box(0); -} -x_252 = lean_array_fget(x_245, x_246); -x_253 = lean_unbox(x_252); -lean_dec(x_252); -x_254 = lean_unsigned_to_nat(1u); -x_255 = lean_nat_add(x_246, x_254); -lean_dec(x_246); -if (lean_is_scalar(x_251)) { - x_256 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_268); + x_275 = lean_box(0); +} +x_276 = lean_array_fget(x_269, x_270); +x_277 = lean_unbox(x_276); +lean_dec(x_276); +x_278 = lean_unsigned_to_nat(1u); +x_279 = lean_nat_add(x_270, x_278); +lean_dec(x_270); +if (lean_is_scalar(x_275)) { + x_280 = lean_alloc_ctor(0, 3, 0); } else { - x_256 = x_251; -} -lean_ctor_set(x_256, 0, x_245); -lean_ctor_set(x_256, 1, x_255); -lean_ctor_set(x_256, 2, x_247); -x_257 = l_Lean_Expr_mvarId_x21(x_22); -x_258 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_257, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_257); -x_259 = lean_ctor_get(x_258, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_258, 1); -lean_inc(x_260); -if (lean_is_exclusive(x_258)) { - lean_ctor_release(x_258, 0); - lean_ctor_release(x_258, 1); - x_261 = x_258; + x_280 = x_275; +} +lean_ctor_set(x_280, 0, x_269); +lean_ctor_set(x_280, 1, x_279); +lean_ctor_set(x_280, 2, x_271); +x_281 = l_Lean_Expr_mvarId_x21(x_22); +x_282 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_281, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_281); +x_283 = lean_ctor_get(x_282, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +if (lean_is_exclusive(x_282)) { + lean_ctor_release(x_282, 0); + lean_ctor_release(x_282, 1); + x_285 = x_282; } else { - lean_dec_ref(x_258); - x_261 = lean_box(0); + lean_dec_ref(x_282); + x_285 = lean_box(0); } -x_262 = l_Lean_BinderInfo_isInstImplicit(x_253); -if (x_262 == 0) +x_286 = l_Lean_BinderInfo_isInstImplicit(x_277); +if (x_286 == 0) { -lean_object* x_263; lean_object* x_264; -lean_dec(x_261); -lean_dec(x_259); +lean_object* x_287; lean_object* x_288; +lean_dec(x_285); +lean_dec(x_283); lean_dec(x_22); lean_inc(x_3); -x_263 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_263, 0, x_3); -lean_ctor_set(x_263, 1, x_256); -x_264 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_264, 0, x_263); -x_23 = x_264; -x_24 = x_260; +x_287 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_287, 0, x_3); +lean_ctor_set(x_287, 1, x_280); +x_288 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_288, 0, x_287); +x_23 = x_288; +x_24 = x_284; goto block_31; } else { -uint8_t x_265; -x_265 = lean_unbox(x_259); -lean_dec(x_259); -if (x_265 == 0) +uint8_t x_289; +x_289 = lean_unbox(x_283); +lean_dec(x_283); +if (x_289 == 0) { -lean_object* x_266; +lean_object* x_290; lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_22); -x_266 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_260); -if (lean_obj_tag(x_266) == 0) -{ -lean_object* x_267; lean_object* x_268; lean_object* x_269; -x_267 = lean_ctor_get(x_266, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_266, 1); -lean_inc(x_268); -lean_dec(x_266); +x_290 = lean_infer_type(x_22, x_15, x_16, x_17, x_18, x_284); +if (lean_obj_tag(x_290) == 0) +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_291 = lean_ctor_get(x_290, 0); +lean_inc(x_291); +x_292 = lean_ctor_get(x_290, 1); +lean_inc(x_292); +lean_dec(x_290); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -lean_inc(x_267); -x_269 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_267, x_15, x_16, x_17, x_18, x_268); -if (lean_obj_tag(x_269) == 0) +lean_inc(x_291); +x_293 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem_synthesizeInstance(x_22, x_291, x_15, x_16, x_17, x_18, x_292); +if (lean_obj_tag(x_293) == 0) +{ +lean_object* x_294; uint8_t x_295; +x_294 = lean_ctor_get(x_293, 0); +lean_inc(x_294); +x_295 = lean_unbox(x_294); +lean_dec(x_294); +if (x_295 == 0) { -lean_object* x_270; uint8_t x_271; -x_270 = lean_ctor_get(x_269, 0); -lean_inc(x_270); -x_271 = lean_unbox(x_270); -lean_dec(x_270); -if (x_271 == 0) +lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; uint8_t x_300; +x_296 = lean_ctor_get(x_293, 1); +lean_inc(x_296); +lean_dec(x_293); +x_297 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_298 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_297, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_296); +x_299 = lean_ctor_get(x_298, 0); +lean_inc(x_299); +x_300 = lean_unbox(x_299); +lean_dec(x_299); +if (x_300 == 0) { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; uint8_t x_276; -x_272 = lean_ctor_get(x_269, 1); -lean_inc(x_272); -lean_dec(x_269); -x_273 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_274 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_273, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_272); -x_275 = lean_ctor_get(x_274, 0); -lean_inc(x_275); -x_276 = lean_unbox(x_275); -lean_dec(x_275); -if (x_276 == 0) -{ -lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -lean_dec(x_267); -lean_dec(x_261); -x_277 = lean_ctor_get(x_274, 1); -lean_inc(x_277); -lean_dec(x_274); -x_278 = lean_box(0); -x_279 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_256, x_278, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_277); -x_280 = lean_ctor_get(x_279, 0); -lean_inc(x_280); -x_281 = lean_ctor_get(x_279, 1); -lean_inc(x_281); -lean_dec(x_279); -x_23 = x_280; -x_24 = x_281; +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; +lean_dec(x_291); +lean_dec(x_285); +x_301 = lean_ctor_get(x_298, 1); +lean_inc(x_301); +lean_dec(x_298); +x_302 = lean_box(0); +x_303 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_280, x_302, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_301); +x_304 = lean_ctor_get(x_303, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_303, 1); +lean_inc(x_305); +lean_dec(x_303); +x_23 = x_304; +x_24 = x_305; goto block_31; } else { -lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; -x_282 = lean_ctor_get(x_274, 1); -lean_inc(x_282); -if (lean_is_exclusive(x_274)) { - lean_ctor_release(x_274, 0); - lean_ctor_release(x_274, 1); - x_283 = x_274; +lean_object* x_306; lean_object* x_307; lean_object* x_308; +x_306 = lean_ctor_get(x_298, 1); +lean_inc(x_306); +if (lean_is_exclusive(x_298)) { + lean_ctor_release(x_298, 0); + lean_ctor_release(x_298, 1); + x_307 = x_298; } else { - lean_dec_ref(x_274); - x_283 = lean_box(0); + lean_dec_ref(x_298); + x_307 = lean_box(0); } -x_284 = lean_ctor_get(x_1, 5); -lean_inc(x_284); -x_285 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_284, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_282); -if (lean_obj_tag(x_285) == 0) +x_308 = l_Lean_Meta_Grind_updateLastTag(x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_306); +if (lean_obj_tag(x_308) == 0) { -lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; -x_286 = lean_ctor_get(x_285, 0); -lean_inc(x_286); -x_287 = lean_ctor_get(x_285, 1); -lean_inc(x_287); +lean_object* x_309; lean_object* x_310; lean_object* x_311; +x_309 = lean_ctor_get(x_308, 1); +lean_inc(x_309); +lean_dec(x_308); +x_310 = lean_ctor_get(x_1, 5); +lean_inc(x_310); +x_311 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_310, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_309); +if (lean_obj_tag(x_311) == 0) +{ +lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; +x_312 = lean_ctor_get(x_311, 0); +lean_inc(x_312); +x_313 = lean_ctor_get(x_311, 1); +lean_inc(x_313); +lean_dec(x_311); +x_314 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; +if (lean_is_scalar(x_307)) { + x_315 = lean_alloc_ctor(7, 2, 0); +} else { + x_315 = x_307; + lean_ctor_set_tag(x_315, 7); +} +lean_ctor_set(x_315, 0, x_314); +lean_ctor_set(x_315, 1, x_312); +x_316 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +if (lean_is_scalar(x_285)) { + x_317 = lean_alloc_ctor(7, 2, 0); +} else { + x_317 = x_285; + lean_ctor_set_tag(x_317, 7); +} +lean_ctor_set(x_317, 0, x_315); +lean_ctor_set(x_317, 1, x_316); +x_318 = l_Lean_indentExpr(x_291); +x_319 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_319, 0, x_317); +lean_ctor_set(x_319, 1, x_318); +x_320 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_320, 0, x_319); +lean_ctor_set(x_320, 1, x_316); +x_321 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_297, x_320, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_313); +x_322 = lean_ctor_get(x_321, 0); +lean_inc(x_322); +x_323 = lean_ctor_get(x_321, 1); +lean_inc(x_323); +lean_dec(x_321); +x_324 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_280, x_322, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_323); +lean_dec(x_322); +x_325 = lean_ctor_get(x_324, 0); +lean_inc(x_325); +x_326 = lean_ctor_get(x_324, 1); +lean_inc(x_326); +lean_dec(x_324); +x_23 = x_325; +x_24 = x_326; +goto block_31; +} +else +{ +lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +lean_dec(x_307); +lean_dec(x_291); lean_dec(x_285); -x_288 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2; -if (lean_is_scalar(x_283)) { - x_289 = lean_alloc_ctor(7, 2, 0); +lean_dec(x_280); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_3); +lean_dec(x_1); +x_327 = lean_ctor_get(x_311, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_311, 1); +lean_inc(x_328); +if (lean_is_exclusive(x_311)) { + lean_ctor_release(x_311, 0); + lean_ctor_release(x_311, 1); + x_329 = x_311; } else { - x_289 = x_283; - lean_ctor_set_tag(x_289, 7); + lean_dec_ref(x_311); + x_329 = lean_box(0); } -lean_ctor_set(x_289, 0, x_288); -lean_ctor_set(x_289, 1, x_286); -x_290 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -if (lean_is_scalar(x_261)) { - x_291 = lean_alloc_ctor(7, 2, 0); +if (lean_is_scalar(x_329)) { + x_330 = lean_alloc_ctor(1, 2, 0); } else { - x_291 = x_261; - lean_ctor_set_tag(x_291, 7); + x_330 = x_329; +} +lean_ctor_set(x_330, 0, x_327); +lean_ctor_set(x_330, 1, x_328); +return x_330; } -lean_ctor_set(x_291, 0, x_289); -lean_ctor_set(x_291, 1, x_290); -x_292 = l_Lean_indentExpr(x_267); -x_293 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_293, 0, x_291); -lean_ctor_set(x_293, 1, x_292); -x_294 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_294, 0, x_293); -lean_ctor_set(x_294, 1, x_290); -x_295 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_273, x_294, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_287); -x_296 = lean_ctor_get(x_295, 0); -lean_inc(x_296); -x_297 = lean_ctor_get(x_295, 1); -lean_inc(x_297); -lean_dec(x_295); -x_298 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_256, x_296, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_297); -lean_dec(x_296); -x_299 = lean_ctor_get(x_298, 0); -lean_inc(x_299); -x_300 = lean_ctor_get(x_298, 1); -lean_inc(x_300); -lean_dec(x_298); -x_23 = x_299; -x_24 = x_300; -goto block_31; } else { -lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; -lean_dec(x_283); -lean_dec(x_267); -lean_dec(x_261); -lean_dec(x_256); +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_dec(x_307); +lean_dec(x_291); +lean_dec(x_285); +lean_dec(x_280); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_301 = lean_ctor_get(x_285, 0); -lean_inc(x_301); -x_302 = lean_ctor_get(x_285, 1); -lean_inc(x_302); -if (lean_is_exclusive(x_285)) { - lean_ctor_release(x_285, 0); - lean_ctor_release(x_285, 1); - x_303 = x_285; +x_331 = lean_ctor_get(x_308, 0); +lean_inc(x_331); +x_332 = lean_ctor_get(x_308, 1); +lean_inc(x_332); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_333 = x_308; } else { - lean_dec_ref(x_285); - x_303 = lean_box(0); + lean_dec_ref(x_308); + x_333 = lean_box(0); } -if (lean_is_scalar(x_303)) { - x_304 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_333)) { + x_334 = lean_alloc_ctor(1, 2, 0); } else { - x_304 = x_303; + x_334 = x_333; } -lean_ctor_set(x_304, 0, x_301); -lean_ctor_set(x_304, 1, x_302); -return x_304; +lean_ctor_set(x_334, 0, x_331); +lean_ctor_set(x_334, 1, x_332); +return x_334; } } } else { -lean_object* x_305; lean_object* x_306; lean_object* x_307; -lean_dec(x_267); -lean_dec(x_261); -x_305 = lean_ctor_get(x_269, 1); -lean_inc(x_305); -lean_dec(x_269); +lean_object* x_335; lean_object* x_336; lean_object* x_337; +lean_dec(x_291); +lean_dec(x_285); +x_335 = lean_ctor_get(x_293, 1); +lean_inc(x_335); +lean_dec(x_293); lean_inc(x_3); -x_306 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_306, 0, x_3); -lean_ctor_set(x_306, 1, x_256); -x_307 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_307, 0, x_306); -x_23 = x_307; -x_24 = x_305; +x_336 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_336, 0, x_3); +lean_ctor_set(x_336, 1, x_280); +x_337 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_337, 0, x_336); +x_23 = x_337; +x_24 = x_335; goto block_31; } } else { -lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; -lean_dec(x_267); -lean_dec(x_261); -lean_dec(x_256); +lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_291); +lean_dec(x_285); +lean_dec(x_280); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_308 = lean_ctor_get(x_269, 0); -lean_inc(x_308); -x_309 = lean_ctor_get(x_269, 1); -lean_inc(x_309); -if (lean_is_exclusive(x_269)) { - lean_ctor_release(x_269, 0); - lean_ctor_release(x_269, 1); - x_310 = x_269; +x_338 = lean_ctor_get(x_293, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_293, 1); +lean_inc(x_339); +if (lean_is_exclusive(x_293)) { + lean_ctor_release(x_293, 0); + lean_ctor_release(x_293, 1); + x_340 = x_293; } else { - lean_dec_ref(x_269); - x_310 = lean_box(0); + lean_dec_ref(x_293); + x_340 = lean_box(0); } -if (lean_is_scalar(x_310)) { - x_311 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_340)) { + x_341 = lean_alloc_ctor(1, 2, 0); } else { - x_311 = x_310; + x_341 = x_340; } -lean_ctor_set(x_311, 0, x_308); -lean_ctor_set(x_311, 1, x_309); -return x_311; +lean_ctor_set(x_341, 0, x_338); +lean_ctor_set(x_341, 1, x_339); +return x_341; } } else { -lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; -lean_dec(x_261); -lean_dec(x_256); +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_285); +lean_dec(x_280); lean_dec(x_22); lean_dec(x_18); lean_dec(x_17); @@ -6625,41 +9759,41 @@ lean_dec(x_16); lean_dec(x_15); lean_dec(x_3); lean_dec(x_1); -x_312 = lean_ctor_get(x_266, 0); -lean_inc(x_312); -x_313 = lean_ctor_get(x_266, 1); -lean_inc(x_313); -if (lean_is_exclusive(x_266)) { - lean_ctor_release(x_266, 0); - lean_ctor_release(x_266, 1); - x_314 = x_266; +x_342 = lean_ctor_get(x_290, 0); +lean_inc(x_342); +x_343 = lean_ctor_get(x_290, 1); +lean_inc(x_343); +if (lean_is_exclusive(x_290)) { + lean_ctor_release(x_290, 0); + lean_ctor_release(x_290, 1); + x_344 = x_290; } else { - lean_dec_ref(x_266); - x_314 = lean_box(0); + lean_dec_ref(x_290); + x_344 = lean_box(0); } -if (lean_is_scalar(x_314)) { - x_315 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_344)) { + x_345 = lean_alloc_ctor(1, 2, 0); } else { - x_315 = x_314; + x_345 = x_344; } -lean_ctor_set(x_315, 0, x_312); -lean_ctor_set(x_315, 1, x_313); -return x_315; +lean_ctor_set(x_345, 0, x_342); +lean_ctor_set(x_345, 1, x_343); +return x_345; } } else { -lean_object* x_316; lean_object* x_317; -lean_dec(x_261); +lean_object* x_346; lean_object* x_347; +lean_dec(x_285); lean_dec(x_22); lean_inc(x_3); -x_316 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_316, 0, x_3); -lean_ctor_set(x_316, 1, x_256); -x_317 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_317, 0, x_316); -x_23 = x_317; -x_24 = x_260; +x_346 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_346, 0, x_3); +lean_ctor_set(x_346, 1, x_280); +x_347 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_347, 0, x_346); +x_23 = x_347; +x_24 = x_284; goto block_31; } } @@ -7379,8 +10513,6 @@ x_26 = lean_ctor_get(x_3, 5); lean_inc(x_26); lean_dec(x_3); x_27 = lean_ctor_get(x_4, 1); -lean_inc(x_27); -lean_dec(x_4); x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_26, x_24, x_27, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_25); return x_28; } @@ -7391,7 +10523,12 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_4); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_3); x_29 = !lean_is_exclusive(x_23); if (x_29 == 0) @@ -7414,17 +10551,7 @@ return x_32; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -static size_t _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1() { +static size_t _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1() { _start: { lean_object* x_1; size_t x_2; @@ -7433,7 +10560,7 @@ x_2 = lean_array_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -7441,16 +10568,16 @@ x_1 = lean_mk_string_unchecked("failed to instantiate ", 22, 22); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4() { _start: { lean_object* x_1; @@ -7458,149 +10585,153 @@ x_1 = lean_mk_string_unchecked(", failed to instantiate non propositional argume return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { _start: { -lean_object* x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_nat_dec_lt(x_22, x_1); -if (x_23 == 0) +lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; +x_22 = l_Lean_mkAppN(x_1, x_2); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_dec_lt(x_23, x_3); +if (x_24 == 0) { -uint8_t x_292; -x_292 = 1; -x_24 = x_292; -x_25 = x_21; -goto block_291; +uint8_t x_327; +x_327 = 1; +x_25 = x_327; +x_26 = x_21; +goto block_326; } else { -size_t x_293; lean_object* x_294; lean_object* x_295; uint8_t x_296; -x_293 = lean_usize_of_nat(x_1); -x_294 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_3, x_8, x_293, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -x_296 = lean_unbox(x_295); -lean_dec(x_295); -if (x_296 == 0) +size_t x_328; lean_object* x_329; lean_object* x_330; uint8_t x_331; +x_328 = lean_usize_of_nat(x_3); +x_329 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_2, x_8, x_328, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +x_330 = lean_ctor_get(x_329, 0); +lean_inc(x_330); +x_331 = lean_unbox(x_330); +lean_dec(x_330); +if (x_331 == 0) { -lean_object* x_297; uint8_t x_298; -x_297 = lean_ctor_get(x_294, 1); -lean_inc(x_297); -lean_dec(x_294); -x_298 = 1; -x_24 = x_298; -x_25 = x_297; -goto block_291; +lean_object* x_332; uint8_t x_333; +x_332 = lean_ctor_get(x_329, 1); +lean_inc(x_332); +lean_dec(x_329); +x_333 = 1; +x_25 = x_333; +x_26 = x_332; +goto block_326; } else { -lean_object* x_299; uint8_t x_300; -x_299 = lean_ctor_get(x_294, 1); -lean_inc(x_299); -lean_dec(x_294); -x_300 = 0; -x_24 = x_300; -x_25 = x_299; -goto block_291; +lean_object* x_334; uint8_t x_335; +x_334 = lean_ctor_get(x_329, 1); +lean_inc(x_334); +lean_dec(x_329); +x_335 = 0; +x_25 = x_335; +x_26 = x_334; +goto block_326; } } -block_291: +block_326: { -if (x_24 == 0) +if (x_25 == 0) { -lean_object* x_26; -x_26 = l_Lean_mkAppN(x_2, x_3); -if (x_23 == 0) +if (x_24 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_101; size_t x_102; lean_object* x_103; -x_101 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_102 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; +lean_object* x_27; lean_object* x_28; lean_object* x_113; size_t x_114; lean_object* x_115; +x_113 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_114 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1; lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_7); -x_103 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_6, x_101, x_7, x_101, x_102, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); -if (lean_obj_tag(x_103) == 0) +x_115 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_6, x_113, x_7, x_113, x_114, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_26); +if (lean_obj_tag(x_115) == 0) { -lean_object* x_104; lean_object* x_105; -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -lean_dec(x_104); -if (lean_obj_tag(x_105) == 0) +lean_object* x_116; lean_object* x_117; +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +lean_dec(x_116); +if (lean_obj_tag(x_117) == 0) { -lean_object* x_106; -x_106 = lean_ctor_get(x_103, 1); -lean_inc(x_106); -lean_dec(x_103); +lean_object* x_118; +x_118 = lean_ctor_get(x_115, 1); +lean_inc(x_118); +lean_dec(x_115); x_27 = x_9; -x_28 = x_106; -goto block_100; +x_28 = x_118; +goto block_112; } else { -lean_object* x_107; lean_object* x_108; +lean_object* x_119; lean_object* x_120; lean_dec(x_9); -x_107 = lean_ctor_get(x_103, 1); -lean_inc(x_107); -lean_dec(x_103); -x_108 = lean_ctor_get(x_105, 0); -lean_inc(x_108); -lean_dec(x_105); -x_27 = x_108; -x_28 = x_107; -goto block_100; +x_119 = lean_ctor_get(x_115, 1); +lean_inc(x_119); +lean_dec(x_115); +x_120 = lean_ctor_get(x_117, 0); +lean_inc(x_120); +lean_dec(x_117); +x_27 = x_120; +x_28 = x_119; +goto block_112; } } else { -uint8_t x_109; -lean_dec(x_26); +uint8_t x_121; +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_9); -lean_dec(x_5); lean_dec(x_4); -x_109 = !lean_is_exclusive(x_103); -if (x_109 == 0) +x_121 = !lean_is_exclusive(x_115); +if (x_121 == 0) { -return x_103; +return x_115; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_103, 0); -x_111 = lean_ctor_get(x_103, 1); -lean_inc(x_111); -lean_inc(x_110); -lean_dec(x_103); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_115, 0); +x_123 = lean_ctor_get(x_115, 1); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_115); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +return x_124; } } -block_100: +block_112: { if (lean_obj_tag(x_27) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; x_29 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; x_30 = lean_box(0); -x_31 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_29, x_4, x_5, x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_28); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_29, x_4, x_5, x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_28); return x_31; } else @@ -7624,7 +10755,7 @@ lean_inc(x_37); lean_dec(x_34); x_38 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; x_39 = lean_box(0); -x_40 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_38, x_4, x_5, x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_37); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_38, x_4, x_5, x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_37); return x_40; } else @@ -7633,246 +10764,355 @@ uint8_t x_41; x_41 = !lean_is_exclusive(x_34); if (x_41 == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_object* x_42; lean_object* x_43; lean_object* x_44; x_42 = lean_ctor_get(x_34, 1); x_43 = lean_ctor_get(x_34, 0); lean_dec(x_43); -x_44 = lean_ctor_get(x_4, 5); -lean_inc(x_44); -x_45 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_44, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_42); -if (lean_obj_tag(x_45) == 0) +x_44 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_42); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_45, 0); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_ctor_get(x_4, 5); lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); +x_47 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_46, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_45); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -x_48 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_47); -if (lean_obj_tag(x_48) == 0) +x_50 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_49); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_46); -lean_ctor_set(x_34, 0, x_51); -x_52 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_34); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Lean_indentExpr(x_49); +lean_ctor_set(x_34, 1, x_48); +lean_ctor_set(x_34, 0, x_53); +x_54 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 0, x_34); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_56 = l_Lean_indentExpr(x_51); x_57 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_57, 0, x_55); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_57, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_50); -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -x_61 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_62 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_61, x_4, x_5, x_59, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_60); -lean_dec(x_59); -return x_62; +x_58 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_59, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_52); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_64 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_63, x_4, x_5, x_61, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_62); +lean_dec(x_61); +return x_64; } else { -uint8_t x_63; -lean_dec(x_46); +uint8_t x_65; +lean_dec(x_48); +lean_free_object(x_34); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_65 = !lean_is_exclusive(x_50); +if (x_65 == 0) +{ +return x_50; +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_50, 0); +x_67 = lean_ctor_get(x_50, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_50); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; +} +} +} +else +{ +uint8_t x_69; lean_free_object(x_34); -lean_dec(x_26); +lean_dec(x_32); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_63 = !lean_is_exclusive(x_48); -if (x_63 == 0) +x_69 = !lean_is_exclusive(x_47); +if (x_69 == 0) { -return x_48; +return x_47; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_48, 0); -x_65 = lean_ctor_get(x_48, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_48); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_47, 0); +x_71 = lean_ctor_get(x_47, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_47); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } } else { -uint8_t x_67; +uint8_t x_73; lean_free_object(x_34); lean_dec(x_32); -lean_dec(x_26); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_67 = !lean_is_exclusive(x_45); -if (x_67 == 0) +x_73 = !lean_is_exclusive(x_44); +if (x_73 == 0) { -return x_45; +return x_44; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_45, 0); -x_69 = lean_ctor_get(x_45, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_45); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_44, 0); +x_75 = lean_ctor_get(x_44, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_44); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_34, 1); -lean_inc(x_71); +lean_object* x_77; lean_object* x_78; +x_77 = lean_ctor_get(x_34, 1); +lean_inc(x_77); lean_dec(x_34); -x_72 = lean_ctor_get(x_4, 5); -lean_inc(x_72); -x_73 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_72, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_71); -if (lean_obj_tag(x_73) == 0) +x_78 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_77); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_73, 1); -lean_inc(x_75); -lean_dec(x_73); +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = lean_ctor_get(x_4, 5); +lean_inc(x_80); +x_81 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_80, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_79); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -x_76 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_75); -if (lean_obj_tag(x_76) == 0) +x_84 = lean_infer_type(x_32, x_17, x_18, x_19, x_20, x_83); +if (lean_obj_tag(x_84) == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_74); -x_81 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -x_82 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Lean_indentExpr(x_77); -x_84 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -x_85 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_86 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_86, 0, x_84); -lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_86, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_78); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_90 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_91 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_90, x_4, x_5, x_88, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_89); -lean_dec(x_88); -return x_91; +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +x_88 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_82); +x_89 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +x_90 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_indentExpr(x_85); +x_92 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +x_93 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_94 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +x_95 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_33, x_94, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_86); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_99 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_98, x_4, x_5, x_96, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_97); +lean_dec(x_96); +return x_99; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_dec(x_74); -lean_dec(x_26); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_82); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_92 = lean_ctor_get(x_76, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_76, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_94 = x_76; +x_100 = lean_ctor_get(x_84, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_84, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_84)) { + lean_ctor_release(x_84, 0); + lean_ctor_release(x_84, 1); + x_102 = x_84; } else { - lean_dec_ref(x_76); - x_94 = lean_box(0); + lean_dec_ref(x_84); + x_102 = lean_box(0); } -if (lean_is_scalar(x_94)) { - x_95 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_102)) { + x_103 = lean_alloc_ctor(1, 2, 0); } else { - x_95 = x_94; + x_103 = x_102; } -lean_ctor_set(x_95, 0, x_92); -lean_ctor_set(x_95, 1, x_93); -return x_95; +lean_ctor_set(x_103, 0, x_100); +lean_ctor_set(x_103, 1, x_101); +return x_103; } } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_dec(x_32); -lean_dec(x_26); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_96 = lean_ctor_get(x_73, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_73, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_98 = x_73; +x_104 = lean_ctor_get(x_81, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_81, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_106 = x_81; } else { - lean_dec_ref(x_73); - x_98 = lean_box(0); + lean_dec_ref(x_81); + x_106 = lean_box(0); } -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(1, 2, 0); } else { - x_99 = x_98; + x_107 = x_106; } -lean_ctor_set(x_99, 0, x_96); -lean_ctor_set(x_99, 1, x_97); -return x_99; +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_105); +return x_107; +} +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_32); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_108 = lean_ctor_get(x_78, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_78, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_110 = x_78; +} else { + lean_dec_ref(x_78); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; } } } @@ -7881,362 +11121,476 @@ return x_99; } else { -uint8_t x_113; -x_113 = lean_nat_dec_le(x_1, x_1); -if (x_113 == 0) +uint8_t x_125; +x_125 = lean_nat_dec_le(x_3, x_3); +if (x_125 == 0) { -lean_object* x_114; lean_object* x_115; lean_object* x_188; size_t x_189; lean_object* x_190; -x_188 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_189 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1; +lean_object* x_126; lean_object* x_127; lean_object* x_212; size_t x_213; lean_object* x_214; +x_212 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_213 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1; lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_7); -x_190 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_6, x_188, x_7, x_188, x_189, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); -if (lean_obj_tag(x_190) == 0) +x_214 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_6, x_212, x_7, x_212, x_213, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_26); +if (lean_obj_tag(x_214) == 0) { -lean_object* x_191; lean_object* x_192; -x_191 = lean_ctor_get(x_190, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_191, 0); -lean_inc(x_192); -lean_dec(x_191); -if (lean_obj_tag(x_192) == 0) +lean_object* x_215; lean_object* x_216; +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_215, 0); +lean_inc(x_216); +lean_dec(x_215); +if (lean_obj_tag(x_216) == 0) { -lean_object* x_193; -x_193 = lean_ctor_get(x_190, 1); -lean_inc(x_193); -lean_dec(x_190); -x_114 = x_9; -x_115 = x_193; -goto block_187; +lean_object* x_217; +x_217 = lean_ctor_get(x_214, 1); +lean_inc(x_217); +lean_dec(x_214); +x_126 = x_9; +x_127 = x_217; +goto block_211; } else { -lean_object* x_194; lean_object* x_195; +lean_object* x_218; lean_object* x_219; lean_dec(x_9); -x_194 = lean_ctor_get(x_190, 1); -lean_inc(x_194); -lean_dec(x_190); -x_195 = lean_ctor_get(x_192, 0); -lean_inc(x_195); -lean_dec(x_192); -x_114 = x_195; -x_115 = x_194; -goto block_187; +x_218 = lean_ctor_get(x_214, 1); +lean_inc(x_218); +lean_dec(x_214); +x_219 = lean_ctor_get(x_216, 0); +lean_inc(x_219); +lean_dec(x_216); +x_126 = x_219; +x_127 = x_218; +goto block_211; } } else { -uint8_t x_196; -lean_dec(x_26); +uint8_t x_220; +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_9); -lean_dec(x_5); lean_dec(x_4); -x_196 = !lean_is_exclusive(x_190); -if (x_196 == 0) +x_220 = !lean_is_exclusive(x_214); +if (x_220 == 0) { -return x_190; +return x_214; } else { -lean_object* x_197; lean_object* x_198; lean_object* x_199; -x_197 = lean_ctor_get(x_190, 0); -x_198 = lean_ctor_get(x_190, 1); -lean_inc(x_198); -lean_inc(x_197); -lean_dec(x_190); -x_199 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_199, 0, x_197); -lean_ctor_set(x_199, 1, x_198); -return x_199; +lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_221 = lean_ctor_get(x_214, 0); +x_222 = lean_ctor_get(x_214, 1); +lean_inc(x_222); +lean_inc(x_221); +lean_dec(x_214); +x_223 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_223, 0, x_221); +lean_ctor_set(x_223, 1, x_222); +return x_223; } } -block_187: +block_211: { -if (lean_obj_tag(x_114) == 0) +if (lean_obj_tag(x_126) == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_116 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_117 = lean_box(0); -x_118 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_116, x_4, x_5, x_117, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_115); -return x_118; +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_129 = lean_box(0); +x_130 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_128, x_4, x_5, x_129, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_127); +return x_130; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_119 = lean_ctor_get(x_114, 0); -lean_inc(x_119); -lean_dec(x_114); -x_120 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_121 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_120, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_115); -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -x_123 = lean_unbox(x_122); -lean_dec(x_122); -if (x_123 == 0) +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; +x_131 = lean_ctor_get(x_126, 0); +lean_inc(x_131); +lean_dec(x_126); +x_132 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_133 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_132, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_127); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_unbox(x_134); +lean_dec(x_134); +if (x_135 == 0) +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_131); +x_136 = lean_ctor_get(x_133, 1); +lean_inc(x_136); +lean_dec(x_133); +x_137 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_138 = lean_box(0); +x_139 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_137, x_4, x_5, x_138, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_136); +return x_139; +} +else +{ +uint8_t x_140; +x_140 = !lean_is_exclusive(x_133); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_133, 1); +x_142 = lean_ctor_get(x_133, 0); +lean_dec(x_142); +x_143 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_141); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_143, 1); +lean_inc(x_144); +lean_dec(x_143); +x_145 = lean_ctor_get(x_4, 5); +lean_inc(x_145); +x_146 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_145, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_144); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_146, 1); +lean_inc(x_148); +lean_dec(x_146); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_149 = lean_infer_type(x_131, x_17, x_18, x_19, x_20, x_148); +if (lean_obj_tag(x_149) == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +x_150 = lean_ctor_get(x_149, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_149, 1); +lean_inc(x_151); +lean_dec(x_149); +x_152 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +lean_ctor_set_tag(x_133, 7); +lean_ctor_set(x_133, 1, x_147); +lean_ctor_set(x_133, 0, x_152); +x_153 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +x_154 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_154, 0, x_133); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_indentExpr(x_150); +x_156 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +x_157 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_158 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_132, x_158, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_151); +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_159, 1); +lean_inc(x_161); +lean_dec(x_159); +x_162 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_163 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_162, x_4, x_5, x_160, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_161); +lean_dec(x_160); +return x_163; +} +else +{ +uint8_t x_164; +lean_dec(x_147); +lean_free_object(x_133); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_164 = !lean_is_exclusive(x_149); +if (x_164 == 0) { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_119); -x_124 = lean_ctor_get(x_121, 1); -lean_inc(x_124); -lean_dec(x_121); -x_125 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_126 = lean_box(0); -x_127 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_125, x_4, x_5, x_126, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_124); -return x_127; +return x_149; } else { -uint8_t x_128; -x_128 = !lean_is_exclusive(x_121); -if (x_128 == 0) -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_129 = lean_ctor_get(x_121, 1); -x_130 = lean_ctor_get(x_121, 0); -lean_dec(x_130); -x_131 = lean_ctor_get(x_4, 5); -lean_inc(x_131); -x_132 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_131, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_129); -if (lean_obj_tag(x_132) == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_132, 1); -lean_inc(x_134); -lean_dec(x_132); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -x_135 = lean_infer_type(x_119, x_17, x_18, x_19, x_20, x_134); -if (lean_obj_tag(x_135) == 0) -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_135, 1); -lean_inc(x_137); -lean_dec(x_135); -x_138 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -lean_ctor_set_tag(x_121, 7); -lean_ctor_set(x_121, 1, x_133); -lean_ctor_set(x_121, 0, x_138); -x_139 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -x_140 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_140, 0, x_121); -lean_ctor_set(x_140, 1, x_139); -x_141 = l_Lean_indentExpr(x_136); -x_142 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -x_143 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_144 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_144, 0, x_142); -lean_ctor_set(x_144, 1, x_143); -x_145 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_120, x_144, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_137); -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -lean_dec(x_145); -x_148 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_149 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_148, x_4, x_5, x_146, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_147); -lean_dec(x_146); -return x_149; +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_149, 0); +x_166 = lean_ctor_get(x_149, 1); +lean_inc(x_166); +lean_inc(x_165); +lean_dec(x_149); +x_167 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set(x_167, 1, x_166); +return x_167; +} +} } else { -uint8_t x_150; -lean_dec(x_133); -lean_free_object(x_121); -lean_dec(x_26); +uint8_t x_168; +lean_free_object(x_133); +lean_dec(x_131); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_150 = !lean_is_exclusive(x_135); -if (x_150 == 0) +x_168 = !lean_is_exclusive(x_146); +if (x_168 == 0) { -return x_135; +return x_146; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_151 = lean_ctor_get(x_135, 0); -x_152 = lean_ctor_get(x_135, 1); -lean_inc(x_152); -lean_inc(x_151); -lean_dec(x_135); -x_153 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_153, 0, x_151); -lean_ctor_set(x_153, 1, x_152); -return x_153; +lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_169 = lean_ctor_get(x_146, 0); +x_170 = lean_ctor_get(x_146, 1); +lean_inc(x_170); +lean_inc(x_169); +lean_dec(x_146); +x_171 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_171, 0, x_169); +lean_ctor_set(x_171, 1, x_170); +return x_171; } } } else { -uint8_t x_154; -lean_free_object(x_121); -lean_dec(x_119); -lean_dec(x_26); +uint8_t x_172; +lean_free_object(x_133); +lean_dec(x_131); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_154 = !lean_is_exclusive(x_132); -if (x_154 == 0) +x_172 = !lean_is_exclusive(x_143); +if (x_172 == 0) { -return x_132; +return x_143; } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_155 = lean_ctor_get(x_132, 0); -x_156 = lean_ctor_get(x_132, 1); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_132); -x_157 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_157, 0, x_155); -lean_ctor_set(x_157, 1, x_156); -return x_157; +lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_173 = lean_ctor_get(x_143, 0); +x_174 = lean_ctor_get(x_143, 1); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_143); +x_175 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_175, 0, x_173); +lean_ctor_set(x_175, 1, x_174); +return x_175; } } } else { -lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_158 = lean_ctor_get(x_121, 1); -lean_inc(x_158); -lean_dec(x_121); -x_159 = lean_ctor_get(x_4, 5); -lean_inc(x_159); -x_160 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_159, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_158); -if (lean_obj_tag(x_160) == 0) -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_160, 1); -lean_inc(x_162); -lean_dec(x_160); +lean_object* x_176; lean_object* x_177; +x_176 = lean_ctor_get(x_133, 1); +lean_inc(x_176); +lean_dec(x_133); +x_177 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_176); +if (lean_obj_tag(x_177) == 0) +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_178 = lean_ctor_get(x_177, 1); +lean_inc(x_178); +lean_dec(x_177); +x_179 = lean_ctor_get(x_4, 5); +lean_inc(x_179); +x_180 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_179, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_178); +if (lean_obj_tag(x_180) == 0) +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_181 = lean_ctor_get(x_180, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_180, 1); +lean_inc(x_182); +lean_dec(x_180); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -x_163 = lean_infer_type(x_119, x_17, x_18, x_19, x_20, x_162); -if (lean_obj_tag(x_163) == 0) +x_183 = lean_infer_type(x_131, x_17, x_18, x_19, x_20, x_182); +if (lean_obj_tag(x_183) == 0) { -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_164 = lean_ctor_get(x_163, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_163, 1); -lean_inc(x_165); -lean_dec(x_163); -x_166 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -x_167 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_167, 0, x_166); -lean_ctor_set(x_167, 1, x_161); -x_168 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -x_169 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_169, 0, x_167); -lean_ctor_set(x_169, 1, x_168); -x_170 = l_Lean_indentExpr(x_164); -x_171 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -x_172 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_173 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_173, 0, x_171); -lean_ctor_set(x_173, 1, x_172); -x_174 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_120, x_173, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_165); -x_175 = lean_ctor_get(x_174, 0); -lean_inc(x_175); -x_176 = lean_ctor_get(x_174, 1); -lean_inc(x_176); -lean_dec(x_174); -x_177 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_178 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_177, x_4, x_5, x_175, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_176); -lean_dec(x_175); -return x_178; +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_184 = lean_ctor_get(x_183, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_183, 1); +lean_inc(x_185); +lean_dec(x_183); +x_186 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +x_187 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_187, 0, x_186); +lean_ctor_set(x_187, 1, x_181); +x_188 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +x_189 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_189, 0, x_187); +lean_ctor_set(x_189, 1, x_188); +x_190 = l_Lean_indentExpr(x_184); +x_191 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_190); +x_192 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_191); +lean_ctor_set(x_193, 1, x_192); +x_194 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_132, x_193, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_185); +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +lean_dec(x_194); +x_197 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_198 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_197, x_4, x_5, x_195, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_196); +lean_dec(x_195); +return x_198; } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -lean_dec(x_161); -lean_dec(x_26); +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_dec(x_181); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_179 = lean_ctor_get(x_163, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_163, 1); -lean_inc(x_180); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_181 = x_163; +x_199 = lean_ctor_get(x_183, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_183, 1); +lean_inc(x_200); +if (lean_is_exclusive(x_183)) { + lean_ctor_release(x_183, 0); + lean_ctor_release(x_183, 1); + x_201 = x_183; } else { - lean_dec_ref(x_163); - x_181 = lean_box(0); + lean_dec_ref(x_183); + x_201 = lean_box(0); } -if (lean_is_scalar(x_181)) { - x_182 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_201)) { + x_202 = lean_alloc_ctor(1, 2, 0); } else { - x_182 = x_181; + x_202 = x_201; } -lean_ctor_set(x_182, 0, x_179); -lean_ctor_set(x_182, 1, x_180); -return x_182; +lean_ctor_set(x_202, 0, x_199); +lean_ctor_set(x_202, 1, x_200); +return x_202; } } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -lean_dec(x_119); -lean_dec(x_26); +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_dec(x_131); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_183 = lean_ctor_get(x_160, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_160, 1); -lean_inc(x_184); -if (lean_is_exclusive(x_160)) { - lean_ctor_release(x_160, 0); - lean_ctor_release(x_160, 1); - x_185 = x_160; +x_203 = lean_ctor_get(x_180, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_180, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_180)) { + lean_ctor_release(x_180, 0); + lean_ctor_release(x_180, 1); + x_205 = x_180; } else { - lean_dec_ref(x_160); - x_185 = lean_box(0); + lean_dec_ref(x_180); + x_205 = lean_box(0); } -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_205)) { + x_206 = lean_alloc_ctor(1, 2, 0); } else { - x_186 = x_185; + x_206 = x_205; } -lean_ctor_set(x_186, 0, x_183); -lean_ctor_set(x_186, 1, x_184); -return x_186; +lean_ctor_set(x_206, 0, x_203); +lean_ctor_set(x_206, 1, x_204); +return x_206; +} +} +else +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_131); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_207 = lean_ctor_get(x_177, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_177, 1); +lean_inc(x_208); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_209 = x_177; +} else { + lean_dec_ref(x_177); + x_209 = lean_box(0); +} +if (lean_is_scalar(x_209)) { + x_210 = lean_alloc_ctor(1, 2, 0); +} else { + x_210 = x_209; +} +lean_ctor_set(x_210, 0, x_207); +lean_ctor_set(x_210, 1, x_208); +return x_210; } } } @@ -8245,394 +11599,512 @@ return x_186; } else { -size_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; size_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_277; -x_200 = lean_usize_of_nat(x_1); -x_201 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; -x_202 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_3, x_8, x_200, x_201, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); -x_203 = lean_ctor_get(x_202, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_202, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - lean_ctor_release(x_202, 1); - x_205 = x_202; +size_t x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; size_t x_230; lean_object* x_231; lean_object* x_232; lean_object* x_313; +x_224 = lean_usize_of_nat(x_3); +x_225 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2; +x_226 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_2, x_8, x_224, x_225, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_26); +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_226, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_229 = x_226; } else { - lean_dec_ref(x_202); - x_205 = lean_box(0); + lean_dec_ref(x_226); + x_229 = lean_box(0); } -x_206 = lean_array_size(x_203); +x_230 = lean_array_size(x_227); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_7); -x_277 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_6, x_203, x_7, x_203, x_206, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_204); -if (lean_obj_tag(x_277) == 0) +x_313 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_6, x_227, x_7, x_227, x_230, x_8, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_228); +if (lean_obj_tag(x_313) == 0) { -lean_object* x_278; lean_object* x_279; -x_278 = lean_ctor_get(x_277, 0); -lean_inc(x_278); -x_279 = lean_ctor_get(x_278, 0); -lean_inc(x_279); -lean_dec(x_278); -if (lean_obj_tag(x_279) == 0) +lean_object* x_314; lean_object* x_315; +x_314 = lean_ctor_get(x_313, 0); +lean_inc(x_314); +x_315 = lean_ctor_get(x_314, 0); +lean_inc(x_315); +lean_dec(x_314); +if (lean_obj_tag(x_315) == 0) { -lean_object* x_280; -x_280 = lean_ctor_get(x_277, 1); -lean_inc(x_280); -lean_dec(x_277); -x_207 = x_9; -x_208 = x_280; -goto block_276; +lean_object* x_316; +x_316 = lean_ctor_get(x_313, 1); +lean_inc(x_316); +lean_dec(x_313); +x_231 = x_9; +x_232 = x_316; +goto block_312; } else { -lean_object* x_281; lean_object* x_282; +lean_object* x_317; lean_object* x_318; lean_dec(x_9); -x_281 = lean_ctor_get(x_277, 1); -lean_inc(x_281); -lean_dec(x_277); -x_282 = lean_ctor_get(x_279, 0); -lean_inc(x_282); -lean_dec(x_279); -x_207 = x_282; -x_208 = x_281; -goto block_276; +x_317 = lean_ctor_get(x_313, 1); +lean_inc(x_317); +lean_dec(x_313); +x_318 = lean_ctor_get(x_315, 0); +lean_inc(x_318); +lean_dec(x_315); +x_231 = x_318; +x_232 = x_317; +goto block_312; +} +} +else +{ +uint8_t x_319; +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_4); +x_319 = !lean_is_exclusive(x_313); +if (x_319 == 0) +{ +return x_313; +} +else +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; +x_320 = lean_ctor_get(x_313, 0); +x_321 = lean_ctor_get(x_313, 1); +lean_inc(x_321); +lean_inc(x_320); +lean_dec(x_313); +x_322 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_322, 0, x_320); +lean_ctor_set(x_322, 1, x_321); +return x_322; +} +} +block_312: +{ +if (lean_obj_tag(x_231) == 0) +{ +lean_object* x_233; lean_object* x_234; +lean_dec(x_229); +x_233 = lean_box(0); +x_234 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_227, x_4, x_5, x_233, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_232); +lean_dec(x_227); +return x_234; } +else +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; +x_235 = lean_ctor_get(x_231, 0); +lean_inc(x_235); +lean_dec(x_231); +x_236 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_237 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_236, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_232); +x_238 = lean_ctor_get(x_237, 0); +lean_inc(x_238); +x_239 = lean_unbox(x_238); +lean_dec(x_238); +if (x_239 == 0) +{ +lean_object* x_240; lean_object* x_241; lean_object* x_242; +lean_dec(x_235); +lean_dec(x_229); +x_240 = lean_ctor_get(x_237, 1); +lean_inc(x_240); +lean_dec(x_237); +x_241 = lean_box(0); +x_242 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_227, x_4, x_5, x_241, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_240); +lean_dec(x_227); +return x_242; +} +else +{ +uint8_t x_243; +x_243 = !lean_is_exclusive(x_237); +if (x_243 == 0) +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_244 = lean_ctor_get(x_237, 1); +x_245 = lean_ctor_get(x_237, 0); +lean_dec(x_245); +x_246 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_244); +if (lean_obj_tag(x_246) == 0) +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_247 = lean_ctor_get(x_246, 1); +lean_inc(x_247); +lean_dec(x_246); +x_248 = lean_ctor_get(x_4, 5); +lean_inc(x_248); +x_249 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_248, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_247); +if (lean_obj_tag(x_249) == 0) +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_249, 1); +lean_inc(x_251); +lean_dec(x_249); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_252 = lean_infer_type(x_235, x_17, x_18, x_19, x_20, x_251); +if (lean_obj_tag(x_252) == 0) +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; +x_253 = lean_ctor_get(x_252, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_252, 1); +lean_inc(x_254); +lean_dec(x_252); +x_255 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +lean_ctor_set_tag(x_237, 7); +lean_ctor_set(x_237, 1, x_250); +lean_ctor_set(x_237, 0, x_255); +x_256 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +if (lean_is_scalar(x_229)) { + x_257 = lean_alloc_ctor(7, 2, 0); +} else { + x_257 = x_229; + lean_ctor_set_tag(x_257, 7); +} +lean_ctor_set(x_257, 0, x_237); +lean_ctor_set(x_257, 1, x_256); +x_258 = l_Lean_indentExpr(x_253); +x_259 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_259, 0, x_257); +lean_ctor_set(x_259, 1, x_258); +x_260 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_261 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_261, 0, x_259); +lean_ctor_set(x_261, 1, x_260); +x_262 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_236, x_261, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_254); +x_263 = lean_ctor_get(x_262, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_262, 1); +lean_inc(x_264); +lean_dec(x_262); +x_265 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_227, x_4, x_5, x_263, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_264); +lean_dec(x_263); +lean_dec(x_227); +return x_265; } else { -uint8_t x_283; -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +uint8_t x_266; +lean_dec(x_250); +lean_free_object(x_237); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_9); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_283 = !lean_is_exclusive(x_277); -if (x_283 == 0) +x_266 = !lean_is_exclusive(x_252); +if (x_266 == 0) { -return x_277; +return x_252; } else { -lean_object* x_284; lean_object* x_285; lean_object* x_286; -x_284 = lean_ctor_get(x_277, 0); -x_285 = lean_ctor_get(x_277, 1); -lean_inc(x_285); -lean_inc(x_284); -lean_dec(x_277); -x_286 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_286, 0, x_284); -lean_ctor_set(x_286, 1, x_285); -return x_286; -} -} -block_276: -{ -if (lean_obj_tag(x_207) == 0) -{ -lean_object* x_209; lean_object* x_210; -lean_dec(x_205); -x_209 = lean_box(0); -x_210 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_209, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_208); -lean_dec(x_203); -return x_210; +lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_267 = lean_ctor_get(x_252, 0); +x_268 = lean_ctor_get(x_252, 1); +lean_inc(x_268); +lean_inc(x_267); +lean_dec(x_252); +x_269 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_269, 0, x_267); +lean_ctor_set(x_269, 1, x_268); +return x_269; } -else -{ -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; -x_211 = lean_ctor_get(x_207, 0); -lean_inc(x_211); -lean_dec(x_207); -x_212 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_213 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_212, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_208); -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_unbox(x_214); -lean_dec(x_214); -if (x_215 == 0) -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; -lean_dec(x_211); -lean_dec(x_205); -x_216 = lean_ctor_get(x_213, 1); -lean_inc(x_216); -lean_dec(x_213); -x_217 = lean_box(0); -x_218 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_217, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_216); -lean_dec(x_203); -return x_218; } -else -{ -uint8_t x_219; -x_219 = !lean_is_exclusive(x_213); -if (x_219 == 0) -{ -lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_220 = lean_ctor_get(x_213, 1); -x_221 = lean_ctor_get(x_213, 0); -lean_dec(x_221); -x_222 = lean_ctor_get(x_4, 5); -lean_inc(x_222); -x_223 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_222, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_220); -if (lean_obj_tag(x_223) == 0) -{ -lean_object* x_224; lean_object* x_225; lean_object* x_226; -x_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_223, 1); -lean_inc(x_225); -lean_dec(x_223); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -x_226 = lean_infer_type(x_211, x_17, x_18, x_19, x_20, x_225); -if (lean_obj_tag(x_226) == 0) -{ -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_226, 1); -lean_inc(x_228); -lean_dec(x_226); -x_229 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -lean_ctor_set_tag(x_213, 7); -lean_ctor_set(x_213, 1, x_224); -lean_ctor_set(x_213, 0, x_229); -x_230 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -if (lean_is_scalar(x_205)) { - x_231 = lean_alloc_ctor(7, 2, 0); -} else { - x_231 = x_205; - lean_ctor_set_tag(x_231, 7); -} -lean_ctor_set(x_231, 0, x_213); -lean_ctor_set(x_231, 1, x_230); -x_232 = l_Lean_indentExpr(x_227); -x_233 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_233, 0, x_231); -lean_ctor_set(x_233, 1, x_232); -x_234 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_235 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_235, 0, x_233); -lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_212, x_235, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_228); -x_237 = lean_ctor_get(x_236, 0); -lean_inc(x_237); -x_238 = lean_ctor_get(x_236, 1); -lean_inc(x_238); -lean_dec(x_236); -x_239 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_237, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_238); -lean_dec(x_237); -lean_dec(x_203); -return x_239; } else { -uint8_t x_240; -lean_dec(x_224); -lean_free_object(x_213); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +uint8_t x_270; +lean_free_object(x_237); +lean_dec(x_235); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_240 = !lean_is_exclusive(x_226); -if (x_240 == 0) +x_270 = !lean_is_exclusive(x_249); +if (x_270 == 0) { -return x_226; +return x_249; } else { -lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_241 = lean_ctor_get(x_226, 0); -x_242 = lean_ctor_get(x_226, 1); -lean_inc(x_242); -lean_inc(x_241); -lean_dec(x_226); -x_243 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_243, 0, x_241); -lean_ctor_set(x_243, 1, x_242); -return x_243; +lean_object* x_271; lean_object* x_272; lean_object* x_273; +x_271 = lean_ctor_get(x_249, 0); +x_272 = lean_ctor_get(x_249, 1); +lean_inc(x_272); +lean_inc(x_271); +lean_dec(x_249); +x_273 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_273, 0, x_271); +lean_ctor_set(x_273, 1, x_272); +return x_273; } } } else { -uint8_t x_244; -lean_free_object(x_213); -lean_dec(x_211); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +uint8_t x_274; +lean_free_object(x_237); +lean_dec(x_235); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_244 = !lean_is_exclusive(x_223); -if (x_244 == 0) +x_274 = !lean_is_exclusive(x_246); +if (x_274 == 0) { -return x_223; +return x_246; } else { -lean_object* x_245; lean_object* x_246; lean_object* x_247; -x_245 = lean_ctor_get(x_223, 0); -x_246 = lean_ctor_get(x_223, 1); -lean_inc(x_246); -lean_inc(x_245); -lean_dec(x_223); -x_247 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_247, 0, x_245); -lean_ctor_set(x_247, 1, x_246); -return x_247; +lean_object* x_275; lean_object* x_276; lean_object* x_277; +x_275 = lean_ctor_get(x_246, 0); +x_276 = lean_ctor_get(x_246, 1); +lean_inc(x_276); +lean_inc(x_275); +lean_dec(x_246); +x_277 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_277, 0, x_275); +lean_ctor_set(x_277, 1, x_276); +return x_277; } } } else { -lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_248 = lean_ctor_get(x_213, 1); -lean_inc(x_248); -lean_dec(x_213); -x_249 = lean_ctor_get(x_4, 5); -lean_inc(x_249); -x_250 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_249, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_248); -if (lean_obj_tag(x_250) == 0) +lean_object* x_278; lean_object* x_279; +x_278 = lean_ctor_get(x_237, 1); +lean_inc(x_278); +lean_dec(x_237); +x_279 = l_Lean_Meta_Grind_updateLastTag(x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_278); +if (lean_obj_tag(x_279) == 0) { -lean_object* x_251; lean_object* x_252; lean_object* x_253; -x_251 = lean_ctor_get(x_250, 0); -lean_inc(x_251); -x_252 = lean_ctor_get(x_250, 1); -lean_inc(x_252); -lean_dec(x_250); +lean_object* x_280; lean_object* x_281; lean_object* x_282; +x_280 = lean_ctor_get(x_279, 1); +lean_inc(x_280); +lean_dec(x_279); +x_281 = lean_ctor_get(x_4, 5); +lean_inc(x_281); +x_282 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_281, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_280); +if (lean_obj_tag(x_282) == 0) +{ +lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_283 = lean_ctor_get(x_282, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +lean_dec(x_282); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -x_253 = lean_infer_type(x_211, x_17, x_18, x_19, x_20, x_252); -if (lean_obj_tag(x_253) == 0) +x_285 = lean_infer_type(x_235, x_17, x_18, x_19, x_20, x_284); +if (lean_obj_tag(x_285) == 0) { -lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; -x_254 = lean_ctor_get(x_253, 0); -lean_inc(x_254); -x_255 = lean_ctor_get(x_253, 1); -lean_inc(x_255); -lean_dec(x_253); -x_256 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3; -x_257 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_257, 0, x_256); -lean_ctor_set(x_257, 1, x_251); -x_258 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5; -if (lean_is_scalar(x_205)) { - x_259 = lean_alloc_ctor(7, 2, 0); +lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_285, 1); +lean_inc(x_287); +lean_dec(x_285); +x_288 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3; +x_289 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_289, 0, x_288); +lean_ctor_set(x_289, 1, x_283); +x_290 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5; +if (lean_is_scalar(x_229)) { + x_291 = lean_alloc_ctor(7, 2, 0); } else { - x_259 = x_205; - lean_ctor_set_tag(x_259, 7); + x_291 = x_229; + lean_ctor_set_tag(x_291, 7); } -lean_ctor_set(x_259, 0, x_257); -lean_ctor_set(x_259, 1, x_258); -x_260 = l_Lean_indentExpr(x_254); -x_261 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_261, 0, x_259); -lean_ctor_set(x_261, 1, x_260); -x_262 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_263 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_263, 0, x_261); -lean_ctor_set(x_263, 1, x_262); -x_264 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_212, x_263, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_255); -x_265 = lean_ctor_get(x_264, 0); -lean_inc(x_265); -x_266 = lean_ctor_get(x_264, 1); -lean_inc(x_266); -lean_dec(x_264); -x_267 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_26, x_203, x_4, x_5, x_265, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_266); -lean_dec(x_265); -lean_dec(x_203); -return x_267; +lean_ctor_set(x_291, 0, x_289); +lean_ctor_set(x_291, 1, x_290); +x_292 = l_Lean_indentExpr(x_286); +x_293 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_293, 0, x_291); +lean_ctor_set(x_293, 1, x_292); +x_294 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_295 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_295, 0, x_293); +lean_ctor_set(x_295, 1, x_294); +x_296 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_236, x_295, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_287); +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +x_298 = lean_ctor_get(x_296, 1); +lean_inc(x_298); +lean_dec(x_296); +x_299 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_22, x_227, x_4, x_5, x_297, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_298); +lean_dec(x_297); +lean_dec(x_227); +return x_299; } else { -lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -lean_dec(x_251); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; +lean_dec(x_283); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_268 = lean_ctor_get(x_253, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_253, 1); -lean_inc(x_269); -if (lean_is_exclusive(x_253)) { - lean_ctor_release(x_253, 0); - lean_ctor_release(x_253, 1); - x_270 = x_253; +x_300 = lean_ctor_get(x_285, 0); +lean_inc(x_300); +x_301 = lean_ctor_get(x_285, 1); +lean_inc(x_301); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_302 = x_285; } else { - lean_dec_ref(x_253); - x_270 = lean_box(0); + lean_dec_ref(x_285); + x_302 = lean_box(0); } -if (lean_is_scalar(x_270)) { - x_271 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_302)) { + x_303 = lean_alloc_ctor(1, 2, 0); } else { - x_271 = x_270; + x_303 = x_302; } -lean_ctor_set(x_271, 0, x_268); -lean_ctor_set(x_271, 1, x_269); -return x_271; +lean_ctor_set(x_303, 0, x_300); +lean_ctor_set(x_303, 1, x_301); +return x_303; } } else { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -lean_dec(x_211); -lean_dec(x_205); -lean_dec(x_203); -lean_dec(x_26); +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; +lean_dec(x_235); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); -lean_dec(x_5); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_4); -x_272 = lean_ctor_get(x_250, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_250, 1); -lean_inc(x_273); -if (lean_is_exclusive(x_250)) { - lean_ctor_release(x_250, 0); - lean_ctor_release(x_250, 1); - x_274 = x_250; +x_304 = lean_ctor_get(x_282, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_282, 1); +lean_inc(x_305); +if (lean_is_exclusive(x_282)) { + lean_ctor_release(x_282, 0); + lean_ctor_release(x_282, 1); + x_306 = x_282; +} else { + lean_dec_ref(x_282); + x_306 = lean_box(0); +} +if (lean_is_scalar(x_306)) { + x_307 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_250); - x_274 = lean_box(0); + x_307 = x_306; } -if (lean_is_scalar(x_274)) { - x_275 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_307, 0, x_304); +lean_ctor_set(x_307, 1, x_305); +return x_307; +} +} +else +{ +lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; +lean_dec(x_235); +lean_dec(x_229); +lean_dec(x_227); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_4); +x_308 = lean_ctor_get(x_279, 0); +lean_inc(x_308); +x_309 = lean_ctor_get(x_279, 1); +lean_inc(x_309); +if (lean_is_exclusive(x_279)) { + lean_ctor_release(x_279, 0); + lean_ctor_release(x_279, 1); + x_310 = x_279; +} else { + lean_dec_ref(x_279); + x_310 = lean_box(0); +} +if (lean_is_scalar(x_310)) { + x_311 = lean_alloc_ctor(1, 2, 0); } else { - x_275 = x_274; + x_311 = x_310; } -lean_ctor_set(x_275, 0, x_272); -lean_ctor_set(x_275, 1, x_273); -return x_275; +lean_ctor_set(x_311, 0, x_308); +lean_ctor_set(x_311, 1, x_309); +return x_311; } } } @@ -8643,23 +12115,20 @@ return x_275; } else { -lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; +lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_dec(x_9); lean_dec(x_7); -x_287 = lean_ctor_get(x_4, 5); -lean_inc(x_287); +x_323 = lean_ctor_get(x_4, 5); +lean_inc(x_323); lean_dec(x_4); -x_288 = l_Lean_mkAppN(x_2, x_3); -x_289 = lean_ctor_get(x_5, 1); -lean_inc(x_289); -lean_dec(x_5); -x_290 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_287, x_288, x_289, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_25); -return x_290; +x_324 = lean_ctor_get(x_5, 1); +x_325 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance(x_323, x_22, x_324, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_26); +return x_325; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; @@ -8695,7 +12164,7 @@ x_31 = lean_ctor_get(x_28, 1); lean_inc(x_31); lean_dec(x_28); x_32 = lean_box(0); -x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_5, x_6, x_3, x_4, x_7, x_24, x_8, x_27, x_2, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_31); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(x_5, x_3, x_6, x_4, x_7, x_24, x_8, x_27, x_2, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_31); return x_33; } else @@ -8705,9 +12174,14 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); x_34 = !lean_is_exclusive(x_28); @@ -8745,9 +12219,14 @@ lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); x_40 = !lean_is_exclusive(x_28); @@ -8771,7 +12250,7 @@ return x_43; } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8783,7 +12262,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -8796,7 +12275,7 @@ lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_21); lean_ctor_set(x_24, 2, x_23); x_25 = lean_box(0); -x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); @@ -8825,13 +12304,7 @@ x_30 = lean_ctor_get(x_27, 1); lean_inc(x_30); lean_dec(x_27); x_31 = lean_box(0); -x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_6, x_25, x_1, x_2, x_21, x_7, x_8, x_26, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_30); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_6, x_25, x_1, x_2, x_7, x_21, x_8, x_26, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_30); lean_dec(x_21); return x_32; } @@ -8849,7 +12322,6 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); @@ -8895,7 +12367,6 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); @@ -8920,7 +12391,7 @@ return x_42; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; @@ -8931,15 +12402,7 @@ lean_ctor_set(x_14, 1, x_12); return x_14; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1() { _start: { lean_object* x_1; @@ -8947,17 +12410,17 @@ x_1 = lean_mk_string_unchecked("c.assignment.size == numParams\n ", 33, 33); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3() { _start: { lean_object* x_1; @@ -8965,28 +12428,28 @@ x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatch.0.Lean.Me return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4; -x_3 = lean_unsigned_to_nat(172u); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3; +x_3 = lean_unsigned_to_nat(240u); x_4 = lean_unsigned_to_nat(2u); -x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3; +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed), 12, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed), 12, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6() { _start: { lean_object* x_1; @@ -8994,16 +12457,16 @@ x_1 = lean_mk_string_unchecked("unexpected number of parameters at ", 35, 35); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; @@ -9027,9 +12490,8 @@ lean_object* x_22; lean_object* x_23; lean_dec(x_20); lean_dec(x_19); lean_dec(x_17); -lean_dec(x_3); lean_dec(x_1); -x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5; +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4; x_23 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); return x_23; } @@ -9088,7 +12550,6 @@ lean_dec(x_32); lean_dec(x_20); lean_dec(x_19); lean_dec(x_17); -lean_dec(x_3); x_38 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; x_39 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_38, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); x_40 = !lean_is_exclusive(x_39); @@ -9097,54 +12558,96 @@ if (x_40 == 0) lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; x_41 = lean_ctor_get(x_39, 0); x_42 = lean_ctor_get(x_39, 1); -x_43 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; +x_43 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5; x_44 = lean_unbox(x_41); lean_dec(x_41); if (x_44 == 0) { -lean_object* x_45; lean_object* x_46; +lean_object* x_45; lean_object* x_46; +lean_free_object(x_39); +lean_free_object(x_30); +lean_dec(x_1); +x_45 = lean_box(0); +x_46 = lean_apply_12(x_43, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); +return x_46; +} +else +{ +lean_object* x_47; +x_47 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = lean_ctor_get(x_1, 5); +lean_inc(x_49); +lean_dec(x_1); +x_50 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_48); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_51); +lean_ctor_set(x_39, 0, x_53); +x_54 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_ctor_set_tag(x_30, 7); +lean_ctor_set(x_30, 1, x_54); +lean_ctor_set(x_30, 0, x_39); +x_55 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_52); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_apply_12(x_43, x_56, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_57); +return x_58; +} +else +{ +uint8_t x_59; lean_free_object(x_39); lean_free_object(x_30); -lean_dec(x_1); -x_45 = lean_box(0); -x_46 = lean_apply_12(x_43, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); -return x_46; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_59 = !lean_is_exclusive(x_50); +if (x_59 == 0) +{ +return x_50; } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_1, 5); -lean_inc(x_47); -lean_dec(x_1); -x_48 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); -if (lean_obj_tag(x_48) == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; -lean_ctor_set_tag(x_39, 7); -lean_ctor_set(x_39, 1, x_49); -lean_ctor_set(x_39, 0, x_51); -x_52 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -lean_ctor_set_tag(x_30, 7); -lean_ctor_set(x_30, 1, x_52); -lean_ctor_set(x_30, 0, x_39); -x_53 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_50); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_apply_12(x_43, x_54, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_55); -return x_56; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_50, 0); +x_61 = lean_ctor_get(x_50, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_50); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} } else { -uint8_t x_57; +uint8_t x_63; lean_free_object(x_39); lean_free_object(x_30); lean_dec(x_14); @@ -9157,82 +12660,90 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_57 = !lean_is_exclusive(x_48); -if (x_57 == 0) +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_47); +if (x_63 == 0) { -return x_48; +return x_47; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_48, 0); -x_59 = lean_ctor_get(x_48, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_48); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_47, 0); +x_65 = lean_ctor_get(x_47, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_47); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_61 = lean_ctor_get(x_39, 0); -x_62 = lean_ctor_get(x_39, 1); -lean_inc(x_62); -lean_inc(x_61); +lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +x_67 = lean_ctor_get(x_39, 0); +x_68 = lean_ctor_get(x_39, 1); +lean_inc(x_68); +lean_inc(x_67); lean_dec(x_39); -x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; -x_64 = lean_unbox(x_61); -lean_dec(x_61); -if (x_64 == 0) +x_69 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5; +x_70 = lean_unbox(x_67); +lean_dec(x_67); +if (x_70 == 0) { -lean_object* x_65; lean_object* x_66; +lean_object* x_71; lean_object* x_72; lean_free_object(x_30); lean_dec(x_1); -x_65 = lean_box(0); -x_66 = lean_apply_12(x_63, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_62); -return x_66; +x_71 = lean_box(0); +x_72 = lean_apply_12(x_69, x_71, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_68); +return x_72; } else { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_1, 5); -lean_inc(x_67); +lean_object* x_73; +x_73 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_68); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +x_75 = lean_ctor_get(x_1, 5); +lean_inc(x_75); lean_dec(x_1); -x_68 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_67, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_62); -if (lean_obj_tag(x_68) == 0) +x_76 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_75, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_74); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; -x_72 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_69); -x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; lean_ctor_set_tag(x_30, 7); -lean_ctor_set(x_30, 1, x_73); -lean_ctor_set(x_30, 0, x_72); -x_74 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_70); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); -lean_inc(x_76); -lean_dec(x_74); -x_77 = lean_apply_12(x_63, x_75, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_76); -return x_77; +lean_ctor_set(x_30, 1, x_81); +lean_ctor_set(x_30, 0, x_80); +x_82 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_38, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_78); +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +lean_dec(x_82); +x_85 = lean_apply_12(x_69, x_83, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_84); +return x_85; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_free_object(x_30); lean_dec(x_14); lean_dec(x_13); @@ -9244,127 +12755,170 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_78 = lean_ctor_get(x_68, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_68, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_80 = x_68; +x_86 = lean_ctor_get(x_76, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_76, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_88 = x_76; } else { - lean_dec_ref(x_68); - x_80 = lean_box(0); + lean_dec_ref(x_76); + x_88 = lean_box(0); } -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(1, 2, 0); } else { - x_81 = x_80; + x_89 = x_88; } -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_79); -return x_81; +lean_ctor_set(x_89, 0, x_86); +lean_ctor_set(x_89, 1, x_87); +return x_89; +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_free_object(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_90 = lean_ctor_get(x_73, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_73, 1); +lean_inc(x_91); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_92 = x_73; +} else { + lean_dec_ref(x_73); + x_92 = lean_box(0); +} +if (lean_is_scalar(x_92)) { + x_93 = lean_alloc_ctor(1, 2, 0); +} else { + x_93 = x_92; +} +lean_ctor_set(x_93, 0, x_90); +lean_ctor_set(x_93, 1, x_91); +return x_93; } } } } else { -lean_object* x_82; lean_object* x_83; +lean_object* x_94; lean_object* x_95; lean_free_object(x_30); -x_82 = lean_box(0); -x_83 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_32, x_1, x_2, x_19, x_20, x_34, x_17, x_3, x_82, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +x_94 = lean_box(0); +x_95 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_32, x_1, x_2, x_19, x_20, x_34, x_17, x_3, x_94, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); lean_dec(x_20); lean_dec(x_19); lean_dec(x_32); -return x_83; +return x_95; } } else { -lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_84 = lean_ctor_get(x_30, 0); -lean_inc(x_84); +lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_96 = lean_ctor_get(x_30, 0); +lean_inc(x_96); lean_dec(x_30); -x_85 = lean_array_get_size(x_32); -x_86 = lean_nat_dec_eq(x_85, x_19); -lean_dec(x_85); -if (x_86 == 0) +x_97 = lean_array_get_size(x_32); +x_98 = lean_nat_dec_eq(x_97, x_19); +lean_dec(x_97); +if (x_98 == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; -lean_dec(x_84); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +lean_dec(x_96); lean_dec(x_32); lean_dec(x_20); lean_dec(x_19); lean_dec(x_17); -lean_dec(x_3); -x_87 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; -x_88 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); -x_89 = lean_ctor_get(x_88, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_88, 1); -lean_inc(x_90); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_91 = x_88; +x_99 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__2; +x_100 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_99, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_103 = x_100; } else { - lean_dec_ref(x_88); - x_91 = lean_box(0); + lean_dec_ref(x_100); + x_103 = lean_box(0); } -x_92 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6; -x_93 = lean_unbox(x_89); -lean_dec(x_89); -if (x_93 == 0) +x_104 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5; +x_105 = lean_unbox(x_101); +lean_dec(x_101); +if (x_105 == 0) { -lean_object* x_94; lean_object* x_95; -lean_dec(x_91); +lean_object* x_106; lean_object* x_107; +lean_dec(x_103); lean_dec(x_1); -x_94 = lean_box(0); -x_95 = lean_apply_12(x_92, x_94, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_90); -return x_95; +x_106 = lean_box(0); +x_107 = lean_apply_12(x_104, x_106, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_102); +return x_107; } else { -lean_object* x_96; lean_object* x_97; -x_96 = lean_ctor_get(x_1, 5); -lean_inc(x_96); +lean_object* x_108; +x_108 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_102); +if (lean_obj_tag(x_108) == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +lean_dec(x_108); +x_110 = lean_ctor_get(x_1, 5); +lean_inc(x_110); lean_dec(x_1); -x_97 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_96, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_90); -if (lean_obj_tag(x_97) == 0) +x_111 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_110, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_109); +if (lean_obj_tag(x_111) == 0) { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -x_100 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8; -if (lean_is_scalar(x_91)) { - x_101 = lean_alloc_ctor(7, 2, 0); +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7; +if (lean_is_scalar(x_103)) { + x_115 = lean_alloc_ctor(7, 2, 0); } else { - x_101 = x_91; - lean_ctor_set_tag(x_101, 7); + x_115 = x_103; + lean_ctor_set_tag(x_115, 7); } -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_101, 1, x_98); -x_102 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_87, x_103, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_99); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_104, 1); -lean_inc(x_106); -lean_dec(x_104); -x_107 = lean_apply_12(x_92, x_105, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_106); -return x_107; +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_112); +x_116 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_117 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +x_118 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_99, x_117, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_113); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = lean_apply_12(x_104, x_119, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_120); +return x_121; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_91); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_103); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -9375,44 +12929,81 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_108 = lean_ctor_get(x_97, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_97, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_110 = x_97; +x_122 = lean_ctor_get(x_111, 0); +lean_inc(x_122); +x_123 = lean_ctor_get(x_111, 1); +lean_inc(x_123); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_124 = x_111; } else { - lean_dec_ref(x_97); - x_110 = lean_box(0); + lean_dec_ref(x_111); + x_124 = lean_box(0); } -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_124)) { + x_125 = lean_alloc_ctor(1, 2, 0); } else { - x_111 = x_110; + x_125 = x_124; } -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -return x_111; +lean_ctor_set(x_125, 0, x_122); +lean_ctor_set(x_125, 1, x_123); +return x_125; +} +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +lean_dec(x_103); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_126 = lean_ctor_get(x_108, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_108, 1); +lean_inc(x_127); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_128 = x_108; +} else { + lean_dec_ref(x_108); + x_128 = lean_box(0); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(1, 2, 0); +} else { + x_129 = x_128; +} +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_127); +return x_129; } } } else { -lean_object* x_112; lean_object* x_113; -x_112 = lean_box(0); -x_113 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_32, x_1, x_2, x_19, x_20, x_84, x_17, x_3, x_112, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +lean_object* x_130; lean_object* x_131; +x_130 = lean_box(0); +x_131 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_32, x_1, x_2, x_19, x_20, x_96, x_17, x_3, x_130, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); lean_dec(x_20); lean_dec(x_19); lean_dec(x_32); -return x_113; +return x_131; } } } else { -uint8_t x_114; +uint8_t x_132; lean_dec(x_20); lean_dec(x_19); lean_dec(x_17); @@ -9426,31 +13017,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_1); -x_114 = !lean_is_exclusive(x_28); -if (x_114 == 0) +x_132 = !lean_is_exclusive(x_28); +if (x_132 == 0) { return x_28; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_28, 0); -x_116 = lean_ctor_get(x_28, 1); -lean_inc(x_116); -lean_inc(x_115); +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_28, 0); +x_134 = lean_ctor_get(x_28, 1); +lean_inc(x_134); +lean_inc(x_133); lean_dec(x_28); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +return x_135; } } } else { -uint8_t x_118; +uint8_t x_136; lean_dec(x_20); lean_dec(x_19); lean_dec(x_17); @@ -9464,32 +13054,31 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_1); -x_118 = !lean_is_exclusive(x_24); -if (x_118 == 0) +x_136 = !lean_is_exclusive(x_24); +if (x_136 == 0) { return x_24; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_24, 0); -x_120 = lean_ctor_get(x_24, 1); -lean_inc(x_120); -lean_inc(x_119); +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_24, 0); +x_138 = lean_ctor_get(x_24, 1); +lean_inc(x_138); +lean_inc(x_137); lean_dec(x_24); -x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -return x_121; +x_139 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_139, 0, x_137); +lean_ctor_set(x_139, 1, x_138); +return x_139; } } } } else { -uint8_t x_122; +uint8_t x_140; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -9500,30 +13089,29 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_1); -x_122 = !lean_is_exclusive(x_16); -if (x_122 == 0) +x_140 = !lean_is_exclusive(x_16); +if (x_140 == 0) { return x_16; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_123 = lean_ctor_get(x_16, 0); -x_124 = lean_ctor_get(x_16, 1); -lean_inc(x_124); -lean_inc(x_123); +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_16, 0); +x_142 = lean_ctor_get(x_16, 1); +lean_inc(x_142); +lean_inc(x_141); lean_dec(x_16); -x_125 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_125, 0, x_123); -lean_ctor_set(x_125, 1, x_124); -return x_125; +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +return x_143; } } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1() { _start: { lean_object* x_1; @@ -9531,23 +13119,23 @@ x_1 = lean_mk_string_unchecked("assignment", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; x_3 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__3; -x_4 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2; x_17 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__2(x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); @@ -9560,7 +13148,7 @@ x_20 = lean_ctor_get(x_17, 1); lean_inc(x_20); lean_dec(x_17); x_21 = lean_box(0); -x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_20); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_20); lean_dec(x_2); return x_22; } @@ -9570,55 +13158,62 @@ uint8_t x_23; x_23 = !lean_is_exclusive(x_17); if (x_23 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_24; lean_object* x_25; lean_object* x_26; x_24 = lean_ctor_get(x_17, 1); x_25 = lean_ctor_get(x_17, 0); lean_dec(x_25); -x_26 = lean_ctor_get(x_1, 5); -lean_inc(x_26); -x_27 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_24); -if (lean_obj_tag(x_27) == 0) +x_26 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_24); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_28 = lean_ctor_get(x_27, 0); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_ctor_get(x_1, 5); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; +x_29 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_27); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; lean_ctor_set_tag(x_17, 7); -lean_ctor_set(x_17, 1, x_28); -lean_ctor_set(x_17, 0, x_30); -x_31 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_17); -lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_17, 1, x_30); +lean_ctor_set(x_17, 0, x_32); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_17); +lean_ctor_set(x_34, 1, x_33); lean_inc(x_2); -x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); -x_34 = lean_array_to_list(x_33); -x_35 = lean_box(0); -x_36 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_34, x_35); -x_37 = l_Lean_MessageData_ofList(x_36); -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_32); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_30); -x_40 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_39, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_29); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_42); -lean_dec(x_41); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); +x_36 = lean_array_to_list(x_35); +x_37 = lean_box(0); +x_38 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_36, x_37); +x_39 = l_Lean_MessageData_ofList(x_38); +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_32); +x_42 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_43, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_44); +lean_dec(x_43); lean_dec(x_2); -return x_43; +return x_45; } else { -uint8_t x_44; +uint8_t x_46; lean_free_object(x_17); lean_dec(x_14); lean_dec(x_13); @@ -9630,80 +13225,122 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_44 = !lean_is_exclusive(x_27); -if (x_44 == 0) +x_46 = !lean_is_exclusive(x_29); +if (x_46 == 0) { -return x_27; +return x_29; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_27, 0); -x_46 = lean_ctor_get(x_27, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_27); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_29, 0); +x_48 = lean_ctor_get(x_29, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_29); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_17, 1); -lean_inc(x_48); -lean_dec(x_17); -x_49 = lean_ctor_get(x_1, 5); -lean_inc(x_49); -x_50 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_48); -if (lean_obj_tag(x_50) == 0) +uint8_t x_50; +lean_free_object(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_50 = !lean_is_exclusive(x_26); +if (x_50 == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); +return x_26; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_26, 0); +x_52 = lean_ctor_get(x_26, 1); lean_inc(x_52); -lean_dec(x_50); -x_53 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -x_55 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; -x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -lean_inc(x_2); -x_57 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); -x_58 = lean_array_to_list(x_57); -x_59 = lean_box(0); -x_60 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_58, x_59); -x_61 = l_Lean_MessageData_ofList(x_60); +lean_inc(x_51); +lean_dec(x_26); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_17, 1); +lean_inc(x_54); +lean_dec(x_17); +x_55 = l_Lean_Meta_Grind_updateLastTag(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_54); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = lean_ctor_get(x_1, 5); +lean_inc(x_57); +x_58 = l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3(x_57, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_56); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData___spec__1___closed__2; x_62 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_62, 0, x_56); -lean_ctor_set(x_62, 1, x_61); -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_53); -x_64 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_63, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_52); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_66); -lean_dec(x_65); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_59); +x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__6; +x_64 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +lean_inc(x_2); +x_65 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assignmentToMessageData(x_2); +x_66 = lean_array_to_list(x_65); +x_67 = lean_box(0); +x_68 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_66, x_67); +x_69 = l_Lean_MessageData_ofList(x_68); +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_64); +lean_ctor_set(x_70, 1, x_69); +x_71 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_61); +x_72 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7(x_16, x_71, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_60); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_73, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_74); +lean_dec(x_73); lean_dec(x_2); -return x_67; +return x_75; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -9714,35 +13351,71 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_68 = lean_ctor_get(x_50, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_50, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_70 = x_50; +x_76 = lean_ctor_get(x_58, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_58, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_78 = x_58; } else { - lean_dec_ref(x_50); - x_70 = lean_box(0); + lean_dec_ref(x_58); + x_78 = lean_box(0); } -if (lean_is_scalar(x_70)) { - x_71 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); +} else { + x_79 = x_78; +} +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_80 = lean_ctor_get(x_55, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_55, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_82 = x_55; +} else { + lean_dec_ref(x_55); + x_82 = lean_box(0); +} +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(1, 2, 0); } else { - x_71 = x_70; + x_83 = x_82; } -lean_ctor_set(x_71, 0, x_68); -lean_ctor_set(x_71, 1, x_69); -return x_71; +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_81); +return x_83; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; @@ -9805,7 +13478,8 @@ x_26 = lean_ctor_get(x_17, 1); lean_inc(x_26); lean_dec(x_17); x_27 = lean_box(0); -x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(x_14, x_16, x_1, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_14, x_16, x_1, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +lean_dec(x_1); return x_28; } } @@ -9831,7 +13505,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_ _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__9), 13, 1); +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8), 13, 1); lean_closure_set(x_13, 0, x_1); x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1; x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); @@ -10138,11 +13812,85 @@ return x_115; } } } -LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +_start: +{ +lean_object* x_24; +x_24 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_24; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; -x_12 = l_ReaderT_read___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_13; +x_13 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10152,14 +13900,16 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_12; +lean_dec(x_1); +return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_13; -x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_14; +x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10170,11 +13920,10 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_13; +return x_14; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -10194,28 +13943,116 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; _start: { -lean_object* x_24; -x_24 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); -lean_dec(x_8); +size_t x_20; size_t x_21; lean_object* x_22; +x_20 = lean_unbox_usize(x_6); lean_dec(x_6); +x_21 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(x_1, x_2, x_3, x_4, x_5, x_20, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -return x_24; +return x_22; } } -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -lean_object* x_13; -x_13 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_17 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_18 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10223,18 +14060,59 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_1); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +_start: +{ +size_t x_19; size_t x_20; lean_object* x_21; +x_19 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_20 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_13; +return x_21; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__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_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; -x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_16 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_17 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_1, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -10244,12 +14122,68 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_2); +lean_dec(x_2); +x_15 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +_start: +{ +size_t x_22; lean_object* x_23; +x_22 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -return x_14; +return x_23; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -10269,27 +14203,19 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -size_t x_20; size_t x_21; lean_object* x_22; -x_20 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_21 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5(x_1, x_2, x_3, x_4, x_5, x_20, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_21; +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -return x_22; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +return x_21; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -10308,181 +14234,709 @@ lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -size_t x_19; size_t x_20; lean_object* x_21; -x_19 = lean_unbox_usize(x_5); +lean_object* x_21; +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_5); -x_20 = lean_unbox_usize(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__6(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_4); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_15 = lean_st_ref_take(x_5, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Meta_Grind_EMatch_instInhabitedChoice; +x_19 = l_List_head_x21___rarg(x_18, x_16); +x_20 = l_List_tail_x21___rarg(x_16); +lean_dec(x_16); +x_21 = lean_st_ref_set(x_5, x_20, x_17); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_23 = lean_ctor_get(x_21, 1); +x_24 = lean_ctor_get(x_21, 0); +lean_dec(x_24); +x_25 = lean_ctor_get(x_19, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +x_27 = lean_ctor_get(x_19, 2); +lean_inc(x_27); +x_28 = lean_nat_dec_lt(x_26, x_1); +if (x_28 == 0) +{ +lean_object* x_29; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_19); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_2); +lean_ctor_set(x_21, 0, x_29); +return x_21; +} +else +{ +lean_free_object(x_21); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_30; +lean_dec(x_27); +lean_dec(x_26); +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 0); +lean_dec(x_32); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_2); +lean_ctor_set(x_30, 0, x_33); +return x_30; +} +else +{ +lean_object* x_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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_2); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +return x_36; +} +} +else +{ +uint8_t x_37; +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_30); +if (x_37 == 0) +{ +return x_30; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_30); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +else +{ +uint8_t x_41; +x_41 = !lean_is_exclusive(x_19); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_19, 2); +lean_dec(x_42); +x_43 = lean_ctor_get(x_19, 1); +lean_dec(x_43); +x_44 = lean_ctor_get(x_19, 0); +lean_dec(x_44); +x_45 = lean_ctor_get(x_25, 0); +lean_inc(x_45); +switch (lean_obj_tag(x_45)) { +case 0: +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_ctor_get(x_25, 1); +lean_inc(x_46); +lean_dec(x_25); +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +lean_ctor_set(x_19, 0, x_46); +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_19, x_47, x_48, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_47); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_2); +lean_ctor_set(x_49, 0, x_52); +return x_49; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_49, 1); +lean_inc(x_53); +lean_dec(x_49); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_2); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +} +else +{ +uint8_t x_56; +lean_dec(x_2); +x_56 = !lean_is_exclusive(x_49); +if (x_56 == 0) +{ +return x_49; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_49, 0); +x_58 = lean_ctor_get(x_49, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_49); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +case 1: +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = lean_ctor_get(x_25, 1); +lean_inc(x_60); +lean_dec(x_25); +x_61 = lean_ctor_get(x_45, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_45, 1); +lean_inc(x_62); +x_63 = lean_ctor_get(x_45, 2); +lean_inc(x_63); +lean_dec(x_45); +lean_ctor_set(x_19, 0, x_60); +x_64 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(x_19, x_61, x_62, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_62); +if (lean_obj_tag(x_64) == 0) +{ +uint8_t x_65; +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_64, 0); +lean_dec(x_66); +x_67 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_67, 0, x_2); +lean_ctor_set(x_64, 0, x_67); +return x_64; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_64, 1); +lean_inc(x_68); +lean_dec(x_64); +x_69 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_69, 0, x_2); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +return x_70; +} +} +else +{ +uint8_t x_71; +lean_dec(x_2); +x_71 = !lean_is_exclusive(x_64); +if (x_71 == 0) +{ +return x_64; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_64, 0); +x_73 = lean_ctor_get(x_64, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_64); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +default: +{ +lean_object* x_75; uint8_t x_76; +x_75 = lean_ctor_get(x_25, 1); +lean_inc(x_75); +lean_dec(x_25); +x_76 = !lean_is_exclusive(x_45); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +x_77 = lean_ctor_get(x_45, 0); +lean_ctor_set(x_19, 0, x_75); +x_78 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_19, x_77, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); +lean_dec(x_5); lean_dec(x_4); +if (lean_obj_tag(x_78) == 0) +{ +uint8_t x_79; +x_79 = !lean_is_exclusive(x_78); +if (x_79 == 0) +{ +lean_object* x_80; +x_80 = lean_ctor_get(x_78, 0); +lean_dec(x_80); +lean_ctor_set_tag(x_45, 1); +lean_ctor_set(x_45, 0, x_2); +lean_ctor_set(x_78, 0, x_45); +return x_78; +} +else +{ +lean_object* x_81; lean_object* x_82; +x_81 = lean_ctor_get(x_78, 1); +lean_inc(x_81); +lean_dec(x_78); +lean_ctor_set_tag(x_45, 1); +lean_ctor_set(x_45, 0, x_2); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_45); +lean_ctor_set(x_82, 1, x_81); +return x_82; +} +} +else +{ +uint8_t x_83; +lean_free_object(x_45); lean_dec(x_2); -lean_dec(x_1); -return x_21; +x_83 = !lean_is_exclusive(x_78); +if (x_83 == 0) +{ +return x_78; } +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_78, 0); +x_85 = lean_ctor_get(x_78, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_78); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -_start: +} +} +else { -size_t x_19; size_t x_20; lean_object* x_21; -x_19 = lean_unbox_usize(x_5); +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_45, 0); +lean_inc(x_87); +lean_dec(x_45); +lean_ctor_set(x_19, 0, x_75); +x_88 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_19, x_87, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); -x_20 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__7(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_21; +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_89 = lean_ctor_get(x_88, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_90 = x_88; +} else { + lean_dec_ref(x_88); + x_90 = lean_box(0); } +x_91 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_91, 0, x_2); +if (lean_is_scalar(x_90)) { + x_92 = lean_alloc_ctor(0, 2, 0); +} else { + x_92 = x_90; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_89); +return x_92; +} +else { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_2); +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_dec(x_2); -x_17 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_18 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__8(x_1, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -return x_18; +x_93 = lean_ctor_get(x_88, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_88, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_95 = x_88; +} else { + lean_dec_ref(x_88); + x_95 = lean_box(0); } +if (lean_is_scalar(x_95)) { + x_96 = lean_alloc_ctor(1, 2, 0); +} else { + x_96 = x_95; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -_start: +lean_ctor_set(x_96, 0, x_93); +lean_ctor_set(x_96, 1, x_94); +return x_96; +} +} +} +} +} +else { -size_t x_19; size_t x_20; lean_object* x_21; -x_19 = lean_unbox_usize(x_5); +lean_object* x_97; +lean_dec(x_19); +x_97 = lean_ctor_get(x_25, 0); +lean_inc(x_97); +switch (lean_obj_tag(x_97)) { +case 0: +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_98 = lean_ctor_get(x_25, 1); +lean_inc(x_98); +lean_dec(x_25); +x_99 = lean_ctor_get(x_97, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +x_101 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_101, 0, x_98); +lean_ctor_set(x_101, 1, x_26); +lean_ctor_set(x_101, 2, x_27); +x_102 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_101, x_99, x_100, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); -x_20 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_21 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__9(x_1, x_2, x_3, x_4, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_21; +lean_dec(x_99); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_104 = x_102; +} else { + lean_dec_ref(x_102); + x_104 = lean_box(0); } +x_105 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_105, 0, x_2); +if (lean_is_scalar(x_104)) { + x_106 = lean_alloc_ctor(0, 2, 0); +} else { + x_106 = x_104; } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_103); +return x_106; +} +else { -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_2); +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_dec(x_2); -x_16 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_17 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__10(x_1, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +x_107 = lean_ctor_get(x_102, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_102, 1); +lean_inc(x_108); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_109 = x_102; +} else { + lean_dec_ref(x_102); + x_109 = lean_box(0); +} +if (lean_is_scalar(x_109)) { + x_110 = lean_alloc_ctor(1, 2, 0); +} else { + x_110 = x_109; +} +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_108); +return x_110; +} +} +case 1: +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_111 = lean_ctor_get(x_25, 1); +lean_inc(x_111); +lean_dec(x_25); +x_112 = lean_ctor_get(x_97, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_97, 1); +lean_inc(x_113); +x_114 = lean_ctor_get(x_97, 2); +lean_inc(x_114); +lean_dec(x_97); +x_115 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_115, 0, x_111); +lean_ctor_set(x_115, 1, x_26); +lean_ctor_set(x_115, 2, x_27); +x_116 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(x_115, x_112, x_113, x_114, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_1); -return x_17; +lean_dec(x_113); +if (lean_obj_tag(x_116) == 0) +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_118 = x_116; +} else { + lean_dec_ref(x_116); + x_118 = lean_box(0); } +x_119 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_119, 0, x_2); +if (lean_is_scalar(x_118)) { + x_120 = lean_alloc_ctor(0, 2, 0); +} else { + x_120 = x_118; } -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_117); +return x_120; +} +else { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_2); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_dec(x_2); -x_15 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__12___rarg(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_15; +x_121 = lean_ctor_get(x_116, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_116, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_123 = x_116; +} else { + lean_dec_ref(x_116); + x_123 = lean_box(0); +} +if (lean_is_scalar(x_123)) { + x_124 = lean_alloc_ctor(1, 2, 0); +} else { + x_124 = x_123; } +lean_ctor_set(x_124, 0, x_121); +lean_ctor_set(x_124, 1, x_122); +return x_124; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: +} +default: { -lean_object* x_17; -x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_125 = lean_ctor_get(x_25, 1); +lean_inc(x_125); +lean_dec(x_25); +x_126 = lean_ctor_get(x_97, 0); +lean_inc(x_126); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + x_127 = x_97; +} else { + lean_dec_ref(x_97); + x_127 = lean_box(0); +} +x_128 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_128, 0, x_125); +lean_ctor_set(x_128, 1, x_26); +lean_ctor_set(x_128, 2, x_27); +x_129 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_128, x_126, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_23); lean_dec(x_5); +lean_dec(x_4); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_130 = lean_ctor_get(x_129, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_131 = x_129; +} else { + lean_dec_ref(x_129); + x_131 = lean_box(0); +} +if (lean_is_scalar(x_127)) { + x_132 = lean_alloc_ctor(1, 1, 0); +} else { + x_132 = x_127; + lean_ctor_set_tag(x_132, 1); +} +lean_ctor_set(x_132, 0, x_2); +if (lean_is_scalar(x_131)) { + x_133 = lean_alloc_ctor(0, 2, 0); +} else { + x_133 = x_131; +} +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_130); +return x_133; +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_dec(x_127); lean_dec(x_2); -return x_17; +x_134 = lean_ctor_get(x_129, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_129, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_136 = x_129; +} else { + lean_dec_ref(x_129); + x_136 = lean_box(0); } +if (lean_is_scalar(x_136)) { + x_137 = lean_alloc_ctor(1, 2, 0); +} else { + x_137 = x_136; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_ctor_set(x_137, 0, x_134); +lean_ctor_set(x_137, 1, x_135); +return x_137; +} +} +} +} +} +} +} +else { -lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_138 = lean_ctor_get(x_21, 1); +lean_inc(x_138); +lean_dec(x_21); +x_139 = lean_ctor_get(x_19, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_19, 1); +lean_inc(x_140); +x_141 = lean_ctor_get(x_19, 2); +lean_inc(x_141); +x_142 = lean_nat_dec_lt(x_140, x_1); +if (x_142 == 0) +{ +lean_object* x_143; lean_object* x_144; +lean_dec(x_141); +lean_dec(x_140); +lean_dec(x_139); +lean_dec(x_19); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -10492,200 +14946,362 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_14; -} +x_143 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_143, 0, x_2); +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_138); +return x_144; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -_start: +else { -size_t x_22; lean_object* x_23; -x_22 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_1); -return x_23; +if (lean_obj_tag(x_139) == 0) +{ +lean_object* x_145; +lean_dec(x_141); +lean_dec(x_140); +x_145 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +if (lean_obj_tag(x_145) == 0) +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_146 = lean_ctor_get(x_145, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_147 = x_145; +} else { + lean_dec_ref(x_145); + x_147 = lean_box(0); } +x_148 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_148, 0, x_2); +if (lean_is_scalar(x_147)) { + x_149 = lean_alloc_ctor(0, 2, 0); +} else { + x_149 = x_147; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -_start: +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_146); +return x_149; +} +else { -lean_object* x_21; -x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_3); -return x_21; +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_dec(x_2); +x_150 = lean_ctor_get(x_145, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_152 = x_145; +} else { + lean_dec_ref(x_145); + x_152 = lean_box(0); } +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 2, 0); +} else { + x_153 = x_152; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -_start: +lean_ctor_set(x_153, 0, x_150); +lean_ctor_set(x_153, 1, x_151); +return x_153; +} +} +else { -lean_object* x_21; -x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_9); +lean_object* x_154; lean_object* x_155; +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + lean_ctor_release(x_19, 1); + lean_ctor_release(x_19, 2); + x_154 = x_19; +} else { + lean_dec_ref(x_19); + x_154 = lean_box(0); +} +x_155 = lean_ctor_get(x_139, 0); +lean_inc(x_155); +switch (lean_obj_tag(x_155)) { +case 0: +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_156 = lean_ctor_get(x_139, 1); +lean_inc(x_156); +lean_dec(x_139); +x_157 = lean_ctor_get(x_155, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_155, 1); +lean_inc(x_158); +lean_dec(x_155); +if (lean_is_scalar(x_154)) { + x_159 = lean_alloc_ctor(0, 3, 0); +} else { + x_159 = x_154; +} +lean_ctor_set(x_159, 0, x_156); +lean_ctor_set(x_159, 1, x_140); +lean_ctor_set(x_159, 2, x_141); +x_160 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_159, x_157, x_158, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_21; +lean_dec(x_157); +if (lean_obj_tag(x_160) == 0) +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_161 = lean_ctor_get(x_160, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_162 = x_160; +} else { + lean_dec_ref(x_160); + x_162 = lean_box(0); +} +x_163 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_163, 0, x_2); +if (lean_is_scalar(x_162)) { + x_164 = lean_alloc_ctor(0, 2, 0); +} else { + x_164 = x_162; } +lean_ctor_set(x_164, 0, x_163); +lean_ctor_set(x_164, 1, x_161); +return x_164; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_dec(x_2); -lean_dec(x_1); -return x_13; +x_165 = lean_ctor_get(x_160, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_160, 1); +lean_inc(x_166); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_167 = x_160; +} else { + lean_dec_ref(x_160); + x_167 = lean_box(0); } +if (lean_is_scalar(x_167)) { + x_168 = lean_alloc_ctor(1, 2, 0); +} else { + x_168 = x_167; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_166); +return x_168; +} +} +case 1: { -lean_object* x_16; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_169 = lean_ctor_get(x_139, 1); +lean_inc(x_169); +lean_dec(x_139); +x_170 = lean_ctor_get(x_155, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_155, 1); +lean_inc(x_171); +x_172 = lean_ctor_get(x_155, 2); +lean_inc(x_172); +lean_dec(x_155); +if (lean_is_scalar(x_154)) { + x_173 = lean_alloc_ctor(0, 3, 0); +} else { + x_173 = x_154; +} +lean_ctor_set(x_173, 0, x_169); +lean_ctor_set(x_173, 1, x_140); +lean_ctor_set(x_173, 2, x_141); +x_174 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processOffset(x_173, x_170, x_171, x_172, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +lean_dec(x_5); lean_dec(x_4); +lean_dec(x_171); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_176 = x_174; +} else { + lean_dec_ref(x_174); + x_176 = lean_box(0); +} +x_177 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_177, 0, x_2); +if (lean_is_scalar(x_176)) { + x_178 = lean_alloc_ctor(0, 2, 0); +} else { + x_178 = x_176; +} +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_175); +return x_178; +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_dec(x_2); -return x_16; +x_179 = lean_ctor_get(x_174, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_174, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_181 = x_174; +} else { + lean_dec_ref(x_174); + x_181 = lean_box(0); } +if (lean_is_scalar(x_181)) { + x_182 = lean_alloc_ctor(1, 2, 0); +} else { + x_182 = x_181; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: +lean_ctor_set(x_182, 0, x_179); +lean_ctor_set(x_182, 1, x_180); +return x_182; +} +} +default: { -lean_object* x_16; -x_16 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_4); -return x_16; +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_183 = lean_ctor_get(x_139, 1); +lean_inc(x_183); +lean_dec(x_139); +x_184 = lean_ctor_get(x_155, 0); +lean_inc(x_184); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + x_185 = x_155; +} else { + lean_dec_ref(x_155); + x_185 = lean_box(0); } +if (lean_is_scalar(x_154)) { + x_186 = lean_alloc_ctor(0, 3, 0); +} else { + x_186 = x_154; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +lean_ctor_set(x_186, 0, x_183); +lean_ctor_set(x_186, 1, x_140); +lean_ctor_set(x_186, 2, x_141); +x_187 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_186, x_184, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +lean_dec(x_5); +lean_dec(x_4); +if (lean_obj_tag(x_187) == 0) { -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_13; +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_188 = lean_ctor_get(x_187, 1); +lean_inc(x_188); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_189 = x_187; +} else { + lean_dec_ref(x_187); + x_189 = lean_box(0); } +if (lean_is_scalar(x_185)) { + x_190 = lean_alloc_ctor(1, 1, 0); +} else { + x_190 = x_185; + lean_ctor_set_tag(x_190, 1); } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1() { -_start: +lean_ctor_set(x_190, 0, x_2); +if (lean_is_scalar(x_189)) { + x_191 = lean_alloc_ctor(0, 2, 0); +} else { + x_191 = x_189; +} +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_188); +return x_191; +} +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1___boxed), 12, 0); -return x_1; +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +lean_dec(x_185); +lean_dec(x_2); +x_192 = lean_ctor_get(x_187, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_187, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_194 = x_187; +} else { + lean_dec_ref(x_187); + x_194 = lean_box(0); +} +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); +} else { + x_195 = x_194; +} +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_193); +return x_195; +} +} +} } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_13 = lean_st_ref_take(x_3, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Meta_Grind_EMatch_instInhabitedChoice; -x_17 = l_List_head_x21___rarg(x_16, x_14); -x_18 = l_List_tail_x21___rarg(x_14); -lean_dec(x_14); -x_19 = lean_st_ref_set(x_3, x_18, x_15); -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -x_21 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1; -x_22 = lean_ctor_get(x_17, 0); +lean_object* x_15; uint8_t x_16; +lean_dec(x_3); +x_15 = lean_st_ref_get(x_5, x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = l_List_isEmpty___rarg(x_17); +lean_dec(x_17); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_free_object(x_15); +x_20 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_21 = l_Lean_Core_checkSystem(x_20, x_12, x_13, x_18); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_21, 1); lean_inc(x_22); -if (lean_obj_tag(x_22) == 0) +lean_dec(x_21); +x_23 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) { -lean_object* x_23; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -10694,23 +15310,18 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); lean_inc(x_2); -x_23 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -if (lean_obj_tag(x_23) == 0) +x_28 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(x_1, x_2, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_26); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_apply_12(x_21, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); -return x_26; -} -else +lean_object* x_29; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +if (lean_obj_tag(x_29) == 0) { -uint8_t x_27; +uint8_t x_30; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10719,75 +15330,53 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -x_27 = !lean_is_exclusive(x_23); -if (x_27 == 0) +x_30 = !lean_is_exclusive(x_28); +if (x_30 == 0) { -return x_23; +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_28, 0); +lean_dec(x_31); +x_32 = lean_ctor_get(x_29, 0); +lean_inc(x_32); +lean_dec(x_29); +lean_ctor_set(x_28, 0, x_32); +return x_28; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 0); -x_29 = lean_ctor_get(x_23, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_23); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_28, 1); +lean_inc(x_33); +lean_dec(x_28); +x_34 = lean_ctor_get(x_29, 0); +lean_inc(x_34); +lean_dec(x_29); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; } } else { -lean_object* x_31; -x_31 = lean_ctor_get(x_22, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_17); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_33 = lean_ctor_get(x_17, 0); -lean_dec(x_33); -x_34 = lean_ctor_get(x_22, 1); -lean_inc(x_34); -lean_dec(x_22); -x_35 = lean_ctor_get(x_31, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_31, 1); +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_28, 1); lean_inc(x_36); -lean_dec(x_31); -lean_ctor_set(x_17, 0, x_34); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_17, x_35, x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -lean_dec(x_35); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_apply_12(x_21, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); -return x_40; +lean_dec(x_28); +x_37 = lean_ctor_get(x_29, 0); +lean_inc(x_37); +lean_dec(x_29); +x_3 = x_37; +x_14 = x_36; +goto _start; +} } else { -uint8_t x_41; +uint8_t x_39; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10796,72 +15385,70 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -x_41 = !lean_is_exclusive(x_37); -if (x_41 == 0) +x_39 = !lean_is_exclusive(x_28); +if (x_39 == 0) { -return x_37; +return x_28; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_37, 0); -x_43 = lean_ctor_get(x_37, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_37); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_28, 0); +x_41 = lean_ctor_get(x_28, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_28); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_45 = lean_ctor_get(x_17, 1); -x_46 = lean_ctor_get(x_17, 2); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_17); -x_47 = lean_ctor_get(x_22, 1); -lean_inc(x_47); -lean_dec(x_22); -x_48 = lean_ctor_get(x_31, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_31, 1); -lean_inc(x_49); -lean_dec(x_31); -x_50 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_50, 0, x_47); -lean_ctor_set(x_50, 1, x_45); -lean_ctor_set(x_50, 2, x_46); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_51 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processMatch(x_50, x_48, x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -lean_dec(x_48); -if (lean_obj_tag(x_51) == 0) +uint8_t x_43; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_43 = !lean_is_exclusive(x_23); +if (x_43 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = lean_apply_12(x_21, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_53); -return x_54; +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_23, 0); +lean_dec(x_44); +x_45 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +lean_ctor_set(x_23, 0, x_45); +return x_23; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_23, 1); +lean_inc(x_46); +lean_dec(x_23); +x_47 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +uint8_t x_49; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10870,47 +15457,78 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -x_55 = lean_ctor_get(x_51, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_51, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_57 = x_51; -} else { - lean_dec_ref(x_51); - x_57 = lean_box(0); +x_49 = !lean_is_exclusive(x_21); +if (x_49 == 0) +{ +return x_21; } -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 2, 0); -} else { - x_58 = x_57; +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_21, 0); +x_51 = lean_ctor_get(x_21, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_21); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; } } +else +{ +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_ctor_set(x_15, 0, x_2); +return x_15; +} } else { -uint8_t x_59; -x_59 = !lean_is_exclusive(x_17); -if (x_59 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = lean_ctor_get(x_15, 0); +x_54 = lean_ctor_get(x_15, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_15); +x_55 = l_List_isEmpty___rarg(x_53); +lean_dec(x_53); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; +x_57 = l_Lean_Core_checkSystem(x_56, x_12, x_13, x_54); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_60 = lean_ctor_get(x_17, 0); +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_58); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_unbox(x_60); lean_dec(x_60); -x_61 = lean_ctor_get(x_22, 1); -lean_inc(x_61); -lean_dec(x_22); -x_62 = lean_ctor_get(x_31, 0); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_59, 1); lean_inc(x_62); -lean_dec(x_31); -lean_ctor_set(x_17, 0, x_61); +lean_dec(x_59); +x_63 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -10919,21 +15537,18 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_17, x_62, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -if (lean_obj_tag(x_63) == 0) +lean_inc(x_2); +x_64 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(x_1, x_2, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_62); +if (lean_obj_tag(x_64) == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); +lean_object* x_65; +x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_apply_12(x_21, x_64, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_65); -return x_66; -} -else +if (lean_obj_tag(x_65) == 0) { -uint8_t x_67; +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -10942,69 +15557,48 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -x_67 = !lean_is_exclusive(x_63); -if (x_67 == 0) -{ -return x_63; +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_67 = x_64; +} else { + lean_dec_ref(x_64); + x_67 = lean_box(0); } -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_63, 0); -x_69 = lean_ctor_get(x_63, 1); -lean_inc(x_69); +x_68 = lean_ctor_get(x_65, 0); lean_inc(x_68); -lean_dec(x_63); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; -} +lean_dec(x_65); +if (lean_is_scalar(x_67)) { + x_69 = lean_alloc_ctor(0, 2, 0); +} else { + x_69 = x_67; } +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_66); +return x_69; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_71 = lean_ctor_get(x_17, 1); -x_72 = lean_ctor_get(x_17, 2); -lean_inc(x_72); +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_64, 1); +lean_inc(x_70); +lean_dec(x_64); +x_71 = lean_ctor_get(x_65, 0); lean_inc(x_71); -lean_dec(x_17); -x_73 = lean_ctor_get(x_22, 1); -lean_inc(x_73); -lean_dec(x_22); -x_74 = lean_ctor_get(x_31, 0); -lean_inc(x_74); -lean_dec(x_31); -x_75 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_71); -lean_ctor_set(x_75, 2, x_72); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_76 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue(x_75, x_74, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = lean_apply_12(x_21, x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_78); -return x_79; +lean_dec(x_65); +x_3 = x_71; +x_14 = x_70; +goto _start; +} } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -11013,86 +15607,35 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -x_80 = lean_ctor_get(x_76, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_76, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_82 = x_76; +x_73 = lean_ctor_get(x_64, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_64, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_75 = x_64; } else { - lean_dec_ref(x_76); - x_82 = lean_box(0); + lean_dec_ref(x_64); + x_75 = lean_box(0); } -if (lean_is_scalar(x_82)) { - x_83 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_75)) { + x_76 = lean_alloc_ctor(1, 2, 0); } else { - x_83 = x_82; -} -lean_ctor_set(x_83, 0, x_80); -lean_ctor_set(x_83, 1, x_81); -return x_83; -} -} -} -} -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___boxed), 12, 0); -return x_1; + x_76 = x_75; } +lean_ctor_set(x_76, 0, x_73); +lean_ctor_set(x_76, 1, x_74); +return x_76; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; uint8_t x_13; -x_12 = lean_st_ref_get(x_2, x_11); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -x_16 = l_List_isEmpty___rarg(x_14); -lean_dec(x_14); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -lean_free_object(x_12); -x_17 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; -x_18 = l_Lean_Core_checkSystem(x_17, x_9, x_10, x_15); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_unbox(x_21); -lean_dec(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; -x_25 = lean_box(0); -x_26 = lean_apply_12(x_24, x_25, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_23); -return x_26; } else { -uint8_t x_27; +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11100,36 +15643,34 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_20); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_20, 0); -lean_dec(x_28); -x_29 = lean_box(0); -lean_ctor_set(x_20, 0, x_29); -return x_20; +x_77 = lean_ctor_get(x_59, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_78 = x_59; +} else { + lean_dec_ref(x_59); + x_78 = lean_box(0); } -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_20, 1); -lean_inc(x_30); -lean_dec(x_20); -x_31 = lean_box(0); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -return x_32; +x_79 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___lambda__1___closed__2; +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 2, 0); +} else { + x_80 = x_78; } +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +return x_80; } } else { -uint8_t x_33; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11137,32 +15678,35 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_33 = !lean_is_exclusive(x_18); -if (x_33 == 0) -{ -return x_18; +x_81 = lean_ctor_get(x_57, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_57, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_83 = x_57; +} else { + lean_dec_ref(x_57); + x_83 = lean_box(0); } -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_18, 0); -x_35 = lean_ctor_get(x_18, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_18); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +if (lean_is_scalar(x_83)) { + x_84 = lean_alloc_ctor(1, 2, 0); +} else { + x_84 = x_83; } +lean_ctor_set(x_84, 0, x_81); +lean_ctor_set(x_84, 1, x_82); +return x_84; } } else { -lean_object* x_37; +lean_object* x_85; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11170,158 +15714,133 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_37 = lean_box(0); -lean_ctor_set(x_12, 0, x_37); -return x_12; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_2); +lean_ctor_set(x_85, 1, x_54); +return x_85; } } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_38 = lean_ctor_get(x_12, 0); -x_39 = lean_ctor_get(x_12, 1); -lean_inc(x_39); -lean_inc(x_38); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Lean_Meta_Grind_getMaxGeneration___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); -x_40 = l_List_isEmpty___rarg(x_38); -lean_dec(x_38); -if (x_40 == 0) -{ -lean_object* x_41; lean_object* x_42; -x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___lambda__2___closed__2; -x_42 = l_Lean_Core_checkSystem(x_41, x_9, x_10, x_39); -if (lean_obj_tag(x_42) == 0) +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; +x_16 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1(x_13, x_15, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +lean_dec(x_13); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -lean_dec(x_42); -x_44 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_43); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_unbox(x_45); -lean_dec(x_45); -if (x_46 == 0) +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1; -x_49 = lean_box(0); -x_50 = lean_apply_12(x_48, x_49, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); -return x_50; -} -else +uint8_t x_19; +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_51 = lean_ctor_get(x_44, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - lean_ctor_release(x_44, 1); - x_52 = x_44; -} else { - lean_dec_ref(x_44); - x_52 = lean_box(0); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_16, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_16, 0, x_21); +return x_16; } -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; } -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; } +else +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_16); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_16, 0); +lean_dec(x_26); +x_27 = lean_ctor_get(x_18, 0); +lean_inc(x_27); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_27); +return x_16; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_55 = lean_ctor_get(x_42, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_42, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_57 = x_42; -} else { - lean_dec_ref(x_42); - x_57 = lean_box(0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_16, 1); +lean_inc(x_28); +lean_dec(x_16); +x_29 = lean_ctor_get(x_18, 0); +lean_inc(x_29); +lean_dec(x_18); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; } -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 2, 0); -} else { - x_58 = x_57; } -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; } +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_16); +if (x_31 == 0) +{ +return x_16; } else { -lean_object* x_59; lean_object* x_60; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_59 = lean_box(0); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_39); -return x_60; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_16, 0); +x_33 = lean_ctor_get(x_16, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_16); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_15; +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_3); lean_dec(x_1); -return x_13; +return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_15; +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_1); -return x_13; +return x_15; } } LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { @@ -11336,20 +15855,17 @@ uint8_t x_21; x_21 = !lean_is_exclusive(x_20); if (x_21 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_89; +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_123; x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_20, 1); -x_89 = lean_ctor_get_uint8(x_22, sizeof(void*)*10 + 4); -if (x_89 == 0) +x_123 = lean_ctor_get_uint8(x_22, sizeof(void*)*10 + 4); +if (x_123 == 0) { -lean_object* x_90; uint8_t x_91; -x_90 = lean_ctor_get(x_22, 3); -lean_inc(x_90); -x_91 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_90, x_1); -lean_dec(x_90); -if (x_91 == 0) +uint8_t x_124; +x_124 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_22); +if (x_124 == 0) { -lean_object* x_92; +lean_object* x_125; lean_dec(x_22); lean_dec(x_18); lean_dec(x_17); @@ -11364,31 +15880,31 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_92 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_92, 0, x_5); -lean_ctor_set(x_20, 0, x_92); +x_125 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_125, 0, x_5); +lean_ctor_set(x_20, 0, x_125); return x_20; } else { if (x_6 == 0) { -lean_object* x_93; +lean_object* x_126; lean_free_object(x_20); -x_93 = lean_box(0); -x_24 = x_93; -goto block_88; +x_126 = lean_box(0); +x_24 = x_126; +goto block_122; } else { -lean_object* x_94; uint8_t x_95; -x_94 = lean_ctor_get(x_22, 9); -lean_inc(x_94); -x_95 = lean_nat_dec_eq(x_94, x_7); -lean_dec(x_94); -if (x_95 == 0) +lean_object* x_127; uint8_t x_128; +x_127 = lean_ctor_get(x_22, 9); +lean_inc(x_127); +x_128 = lean_nat_dec_eq(x_127, x_7); +lean_dec(x_127); +if (x_128 == 0) { -lean_object* x_96; +lean_object* x_129; lean_dec(x_22); lean_dec(x_18); lean_dec(x_17); @@ -11403,18 +15919,18 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_96 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_96, 0, x_5); -lean_ctor_set(x_20, 0, x_96); +x_129 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_129, 0, x_5); +lean_ctor_set(x_20, 0, x_129); return x_20; } else { -lean_object* x_97; +lean_object* x_130; lean_free_object(x_20); -x_97 = lean_box(0); -x_24 = x_97; -goto block_88; +x_130 = lean_box(0); +x_24 = x_130; +goto block_122; } } } @@ -11423,22 +15939,22 @@ else { if (x_6 == 0) { -lean_object* x_98; +lean_object* x_131; lean_free_object(x_20); -x_98 = lean_box(0); -x_24 = x_98; -goto block_88; +x_131 = lean_box(0); +x_24 = x_131; +goto block_122; } else { -lean_object* x_99; uint8_t x_100; -x_99 = lean_ctor_get(x_22, 9); -lean_inc(x_99); -x_100 = lean_nat_dec_eq(x_99, x_7); -lean_dec(x_99); -if (x_100 == 0) +lean_object* x_132; uint8_t x_133; +x_132 = lean_ctor_get(x_22, 9); +lean_inc(x_132); +x_133 = lean_nat_dec_eq(x_132, x_7); +lean_dec(x_132); +if (x_133 == 0) { -lean_object* x_101; +lean_object* x_134; lean_dec(x_22); lean_dec(x_18); lean_dec(x_17); @@ -11453,24 +15969,24 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_101 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_101, 0, x_5); -lean_ctor_set(x_20, 0, x_101); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_5); +lean_ctor_set(x_20, 0, x_134); return x_20; } else { -lean_object* x_102; +lean_object* x_135; lean_free_object(x_20); -x_102 = lean_box(0); -x_24 = x_102; -goto block_88; +x_135 = lean_box(0); +x_24 = x_135; +goto block_122; } } } -block_88: +block_122: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_dec(x_24); x_25 = lean_ctor_get(x_22, 8); lean_inc(x_25); @@ -11479,6 +15995,14 @@ x_26 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_26, 0, x_2); lean_ctor_set(x_26, 1, x_25); lean_ctor_set(x_26, 2, x_3); +x_27 = !lean_is_exclusive(x_9); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_9, 1); +lean_dec(x_28); +lean_inc(x_1); +lean_ctor_set(x_9, 1, x_1); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); @@ -11487,16 +16011,17 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_27 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_26, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_23); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_26, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_23); lean_dec(x_1); -if (lean_obj_tag(x_27) == 0) +if (lean_obj_tag(x_29) == 0) { -lean_object* x_28; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -if (lean_obj_tag(x_28) == 0) +lean_object* x_30; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +if (lean_obj_tag(x_30) == 0) { -uint8_t x_29; +uint8_t x_31; +lean_dec(x_9); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11506,263 +16031,454 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -x_29 = !lean_is_exclusive(x_27); -if (x_29 == 0) +x_31 = !lean_is_exclusive(x_29); +if (x_31 == 0) { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_27, 0); -lean_dec(x_30); -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_5); -lean_ctor_set(x_27, 0, x_31); -return x_27; +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 0); +lean_dec(x_32); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_5); +lean_ctor_set(x_29, 0, x_33); +return x_29; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_27, 1); -lean_inc(x_32); -lean_dec(x_27); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_5); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -return x_34; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_29, 1); +lean_inc(x_34); +lean_dec(x_29); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_5); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +return x_36; } } else { -lean_object* x_35; uint8_t x_36; -x_35 = lean_ctor_get(x_27, 1); -lean_inc(x_35); -lean_dec(x_27); -x_36 = !lean_is_exclusive(x_28); -if (x_36 == 0) +lean_object* x_37; uint8_t x_38; +x_37 = lean_ctor_get(x_29, 1); +lean_inc(x_37); +lean_dec(x_29); +x_38 = !lean_is_exclusive(x_30); +if (x_38 == 0) { -lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_37 = lean_ctor_get(x_28, 0); -x_38 = lean_st_ref_take(x_10, x_35); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_30, 0); +x_40 = lean_st_ref_take(x_10, x_37); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_38, 1); -x_41 = lean_ctor_get(x_38, 0); -lean_dec(x_41); -x_42 = lean_box(0); -lean_ctor_set_tag(x_38, 1); -lean_ctor_set(x_38, 1, x_42); -lean_ctor_set(x_38, 0, x_37); -x_43 = lean_st_ref_set(x_10, x_38, x_40); -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_42 = lean_ctor_get(x_40, 1); +x_43 = lean_ctor_get(x_40, 0); lean_dec(x_43); -x_45 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_44); -if (lean_obj_tag(x_45) == 0) +x_44 = lean_box(0); +lean_ctor_set_tag(x_40, 1); +lean_ctor_set(x_40, 1, x_44); +lean_ctor_set(x_40, 0, x_39); +x_45 = lean_st_ref_set(x_10, x_40, x_42); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +x_47 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_46); +if (lean_obj_tag(x_47) == 0) { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) +uint8_t x_48; +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -lean_object* x_47; -x_47 = lean_ctor_get(x_45, 0); +lean_object* x_49; +x_49 = lean_ctor_get(x_47, 0); +lean_dec(x_49); +lean_ctor_set(x_30, 0, x_5); +lean_ctor_set(x_47, 0, x_30); +return x_47; +} +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); lean_dec(x_47); -lean_ctor_set(x_28, 0, x_5); -lean_ctor_set(x_45, 0, x_28); -return x_45; +lean_ctor_set(x_30, 0, x_5); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_30); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +else +{ +uint8_t x_52; +lean_free_object(x_30); +lean_dec(x_5); +x_52 = !lean_is_exclusive(x_47); +if (x_52 == 0) +{ +return x_47; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_47, 0); +x_54 = lean_ctor_get(x_47, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_47); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_56 = lean_ctor_get(x_40, 1); +lean_inc(x_56); +lean_dec(x_40); +x_57 = lean_box(0); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_39); +lean_ctor_set(x_58, 1, x_57); +x_59 = lean_st_ref_set(x_10, x_58, x_56); +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_60); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_63 = x_61; +} else { + lean_dec_ref(x_61); + x_63 = lean_box(0); +} +lean_ctor_set(x_30, 0, x_5); +if (lean_is_scalar(x_63)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_63; +} +lean_ctor_set(x_64, 0, x_30); +lean_ctor_set(x_64, 1, x_62); +return x_64; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_free_object(x_30); +lean_dec(x_5); +x_65 = lean_ctor_get(x_61, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_61, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_67 = x_61; +} else { + lean_dec_ref(x_61); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(1, 2, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; +} +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_69 = lean_ctor_get(x_30, 0); +lean_inc(x_69); +lean_dec(x_30); +x_70 = lean_st_ref_take(x_10, x_37); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_72 = x_70; +} else { + lean_dec_ref(x_70); + x_72 = lean_box(0); +} +x_73 = lean_box(0); +if (lean_is_scalar(x_72)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_72; + lean_ctor_set_tag(x_74, 1); +} +lean_ctor_set(x_74, 0, x_69); +lean_ctor_set(x_74, 1, x_73); +x_75 = lean_st_ref_set(x_10, x_74, x_71); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_76); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_79 = x_77; +} else { + lean_dec_ref(x_77); + x_79 = lean_box(0); +} +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_5); +if (lean_is_scalar(x_79)) { + x_81 = lean_alloc_ctor(0, 2, 0); +} else { + x_81 = x_79; +} +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_78); +return x_81; } else { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_45, 1); -lean_inc(x_48); -lean_dec(x_45); -lean_ctor_set(x_28, 0, x_5); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_28); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_5); +x_82 = lean_ctor_get(x_77, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_77, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_84 = x_77; +} else { + lean_dec_ref(x_77); + x_84 = lean_box(0); +} +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_84; +} +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; +} +} } } else { -uint8_t x_50; -lean_free_object(x_28); +uint8_t x_86; +lean_dec(x_9); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_5); -x_50 = !lean_is_exclusive(x_45); -if (x_50 == 0) +x_86 = !lean_is_exclusive(x_29); +if (x_86 == 0) { -return x_45; +return x_29; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_45, 0); -x_52 = lean_ctor_get(x_45, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_45); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -return x_53; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_29, 0); +x_88 = lean_ctor_get(x_29, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_29); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; } } } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_54 = lean_ctor_get(x_38, 1); -lean_inc(x_54); -lean_dec(x_38); -x_55 = lean_box(0); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_37); -lean_ctor_set(x_56, 1, x_55); -x_57 = lean_st_ref_set(x_10, x_56, x_54); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_58); -if (lean_obj_tag(x_59) == 0) +uint8_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get_uint8(x_9, sizeof(void*)*2); +x_91 = lean_ctor_get(x_9, 0); +lean_inc(x_91); +lean_dec(x_9); +lean_inc(x_1); +x_92 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_1); +lean_ctor_set_uint8(x_92, sizeof(void*)*2, x_90); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_93 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_26, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_23); +lean_dec(x_1); +if (lean_obj_tag(x_93) == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_61 = x_59; +lean_object* x_94; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +if (lean_obj_tag(x_94) == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_92); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_96 = x_93; } else { - lean_dec_ref(x_59); - x_61 = lean_box(0); + lean_dec_ref(x_93); + x_96 = lean_box(0); } -lean_ctor_set(x_28, 0, x_5); -if (lean_is_scalar(x_61)) { - x_62 = lean_alloc_ctor(0, 2, 0); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_5); +if (lean_is_scalar(x_96)) { + x_98 = lean_alloc_ctor(0, 2, 0); } else { - x_62 = x_61; + x_98 = x_96; } -lean_ctor_set(x_62, 0, x_28); -lean_ctor_set(x_62, 1, x_60); -return x_62; +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_95); +return x_98; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -lean_free_object(x_28); -lean_dec(x_5); -x_63 = lean_ctor_get(x_59, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_59, 1); -lean_inc(x_64); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_65 = x_59; +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_99 = lean_ctor_get(x_93, 1); +lean_inc(x_99); +lean_dec(x_93); +x_100 = lean_ctor_get(x_94, 0); +lean_inc(x_100); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + x_101 = x_94; } else { - lean_dec_ref(x_59); - x_65 = lean_box(0); + lean_dec_ref(x_94); + x_101 = lean_box(0); } -if (lean_is_scalar(x_65)) { - x_66 = lean_alloc_ctor(1, 2, 0); +x_102 = lean_st_ref_take(x_10, x_99); +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_104 = x_102; } else { - x_66 = x_65; + lean_dec_ref(x_102); + x_104 = lean_box(0); } -lean_ctor_set(x_66, 0, x_63); -lean_ctor_set(x_66, 1, x_64); -return x_66; -} -} -} -else -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_67 = lean_ctor_get(x_28, 0); -lean_inc(x_67); -lean_dec(x_28); -x_68 = lean_st_ref_take(x_10, x_35); -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_70 = x_68; +x_105 = lean_box(0); +if (lean_is_scalar(x_104)) { + x_106 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_68); - x_70 = lean_box(0); + x_106 = x_104; + lean_ctor_set_tag(x_106, 1); } -x_71 = lean_box(0); -if (lean_is_scalar(x_70)) { - x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_100); +lean_ctor_set(x_106, 1, x_105); +x_107 = lean_st_ref_set(x_10, x_106, x_103); +x_108 = lean_ctor_get(x_107, 1); +lean_inc(x_108); +lean_dec(x_107); +x_109 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_92, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_108); +if (lean_obj_tag(x_109) == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_110 = lean_ctor_get(x_109, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_111 = x_109; } else { - x_72 = x_70; - lean_ctor_set_tag(x_72, 1); + lean_dec_ref(x_109); + x_111 = lean_box(0); } -lean_ctor_set(x_72, 0, x_67); -lean_ctor_set(x_72, 1, x_71); -x_73 = lean_st_ref_set(x_10, x_72, x_69); -x_74 = lean_ctor_get(x_73, 1); -lean_inc(x_74); -lean_dec(x_73); -x_75 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_74); -if (lean_obj_tag(x_75) == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_77 = x_75; +if (lean_is_scalar(x_101)) { + x_112 = lean_alloc_ctor(1, 1, 0); } else { - lean_dec_ref(x_75); - x_77 = lean_box(0); + x_112 = x_101; } -x_78 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_78, 0, x_5); -if (lean_is_scalar(x_77)) { - x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_5); +if (lean_is_scalar(x_111)) { + x_113 = lean_alloc_ctor(0, 2, 0); } else { - x_79 = x_77; + x_113 = x_111; } -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_76); -return x_79; +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_110); +return x_113; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_101); lean_dec(x_5); -x_80 = lean_ctor_get(x_75, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_75, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_82 = x_75; +x_114 = lean_ctor_get(x_109, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_109, 1); +lean_inc(x_115); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_116 = x_109; } else { - lean_dec_ref(x_75); - x_82 = lean_box(0); + lean_dec_ref(x_109); + x_116 = lean_box(0); } -if (lean_is_scalar(x_82)) { - x_83 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_116)) { + x_117 = lean_alloc_ctor(1, 2, 0); } else { - x_83 = x_82; -} -lean_ctor_set(x_83, 0, x_80); -lean_ctor_set(x_83, 1, x_81); -return x_83; + x_117 = x_116; } +lean_ctor_set(x_117, 0, x_114); +lean_ctor_set(x_117, 1, x_115); +return x_117; } } } else { -uint8_t x_84; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_92); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11772,49 +16488,48 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_5); -x_84 = !lean_is_exclusive(x_27); -if (x_84 == 0) -{ -return x_27; +x_118 = lean_ctor_get(x_93, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_93, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_120 = x_93; +} else { + lean_dec_ref(x_93); + x_120 = lean_box(0); } -else -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_27, 0); -x_86 = lean_ctor_get(x_27, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_27); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); +} else { + x_121 = x_120; +} +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; } } } } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_138; -x_103 = lean_ctor_get(x_20, 0); -x_104 = lean_ctor_get(x_20, 1); -lean_inc(x_104); -lean_inc(x_103); +lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_175; +x_136 = lean_ctor_get(x_20, 0); +x_137 = lean_ctor_get(x_20, 1); +lean_inc(x_137); +lean_inc(x_136); lean_dec(x_20); -x_138 = lean_ctor_get_uint8(x_103, sizeof(void*)*10 + 4); -if (x_138 == 0) +x_175 = lean_ctor_get_uint8(x_136, sizeof(void*)*10 + 4); +if (x_175 == 0) { -lean_object* x_139; uint8_t x_140; -x_139 = lean_ctor_get(x_103, 3); -lean_inc(x_139); -x_140 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_139, x_1); -lean_dec(x_139); -if (x_140 == 0) +uint8_t x_176; +x_176 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_136); +if (x_176 == 0) { -lean_object* x_141; lean_object* x_142; -lean_dec(x_103); +lean_object* x_177; lean_object* x_178; +lean_dec(x_136); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11828,33 +16543,33 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_141 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_141, 0, x_5); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_104); -return x_142; +x_177 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_177, 0, x_5); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_137); +return x_178; } else { if (x_6 == 0) { -lean_object* x_143; -x_143 = lean_box(0); -x_105 = x_143; -goto block_137; +lean_object* x_179; +x_179 = lean_box(0); +x_138 = x_179; +goto block_174; } else { -lean_object* x_144; uint8_t x_145; -x_144 = lean_ctor_get(x_103, 9); -lean_inc(x_144); -x_145 = lean_nat_dec_eq(x_144, x_7); -lean_dec(x_144); -if (x_145 == 0) +lean_object* x_180; uint8_t x_181; +x_180 = lean_ctor_get(x_136, 9); +lean_inc(x_180); +x_181 = lean_nat_dec_eq(x_180, x_7); +lean_dec(x_180); +if (x_181 == 0) { -lean_object* x_146; lean_object* x_147; -lean_dec(x_103); +lean_object* x_182; lean_object* x_183; +lean_dec(x_136); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11868,19 +16583,19 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_146 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_146, 0, x_5); -x_147 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_147, 0, x_146); -lean_ctor_set(x_147, 1, x_104); -return x_147; +x_182 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_182, 0, x_5); +x_183 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_137); +return x_183; } else { -lean_object* x_148; -x_148 = lean_box(0); -x_105 = x_148; -goto block_137; +lean_object* x_184; +x_184 = lean_box(0); +x_138 = x_184; +goto block_174; } } } @@ -11889,22 +16604,22 @@ else { if (x_6 == 0) { -lean_object* x_149; -x_149 = lean_box(0); -x_105 = x_149; -goto block_137; +lean_object* x_185; +x_185 = lean_box(0); +x_138 = x_185; +goto block_174; } else { -lean_object* x_150; uint8_t x_151; -x_150 = lean_ctor_get(x_103, 9); -lean_inc(x_150); -x_151 = lean_nat_dec_eq(x_150, x_7); -lean_dec(x_150); -if (x_151 == 0) +lean_object* x_186; uint8_t x_187; +x_186 = lean_ctor_get(x_136, 9); +lean_inc(x_186); +x_187 = lean_nat_dec_eq(x_186, x_7); +lean_dec(x_186); +if (x_187 == 0) { -lean_object* x_152; lean_object* x_153; -lean_dec(x_103); +lean_object* x_188; lean_object* x_189; +lean_dec(x_136); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11918,33 +16633,53 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_152 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_152, 0, x_5); -x_153 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_104); -return x_153; +x_188 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_188, 0, x_5); +x_189 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_189, 0, x_188); +lean_ctor_set(x_189, 1, x_137); +return x_189; } else { -lean_object* x_154; -x_154 = lean_box(0); -x_105 = x_154; -goto block_137; +lean_object* x_190; +x_190 = lean_box(0); +x_138 = x_190; +goto block_174; } } } -block_137: +block_174: { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -lean_dec(x_105); -x_106 = lean_ctor_get(x_103, 8); -lean_inc(x_106); -lean_dec(x_103); -x_107 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_107, 0, x_2); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_3); +lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_138); +x_139 = lean_ctor_get(x_136, 8); +lean_inc(x_139); +lean_dec(x_136); +x_140 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_140, 0, x_2); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_3); +x_141 = lean_ctor_get_uint8(x_9, sizeof(void*)*2); +x_142 = lean_ctor_get(x_9, 0); +lean_inc(x_142); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + x_143 = x_9; +} else { + lean_dec_ref(x_9); + x_143 = lean_box(0); +} +lean_inc(x_1); +if (lean_is_scalar(x_143)) { + x_144 = lean_alloc_ctor(0, 2, 1); +} else { + x_144 = x_143; +} +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_1); +lean_ctor_set_uint8(x_144, sizeof(void*)*2, x_141); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); @@ -11953,16 +16688,17 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_108 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_107, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_104); +x_145 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArgs_x3f(x_140, x_4, x_1, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_137); lean_dec(x_1); -if (lean_obj_tag(x_108) == 0) +if (lean_obj_tag(x_145) == 0) { -lean_object* x_109; -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -if (lean_obj_tag(x_109) == 0) +lean_object* x_146; +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_144); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -11972,127 +16708,127 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - lean_ctor_release(x_108, 1); - x_111 = x_108; +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_148 = x_145; } else { - lean_dec_ref(x_108); - x_111 = lean_box(0); + lean_dec_ref(x_145); + x_148 = lean_box(0); } -x_112 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_112, 0, x_5); -if (lean_is_scalar(x_111)) { - x_113 = lean_alloc_ctor(0, 2, 0); +x_149 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_149, 0, x_5); +if (lean_is_scalar(x_148)) { + x_150 = lean_alloc_ctor(0, 2, 0); } else { - x_113 = x_111; + x_150 = x_148; } -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_110); -return x_113; +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_147); +return x_150; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_114 = lean_ctor_get(x_108, 1); -lean_inc(x_114); -lean_dec(x_108); -x_115 = lean_ctor_get(x_109, 0); -lean_inc(x_115); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - x_116 = x_109; +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +lean_dec(x_145); +x_152 = lean_ctor_get(x_146, 0); +lean_inc(x_152); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + x_153 = x_146; } else { - lean_dec_ref(x_109); - x_116 = lean_box(0); + lean_dec_ref(x_146); + x_153 = lean_box(0); } -x_117 = lean_st_ref_take(x_10, x_114); -x_118 = lean_ctor_get(x_117, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_119 = x_117; +x_154 = lean_st_ref_take(x_10, x_151); +x_155 = lean_ctor_get(x_154, 1); +lean_inc(x_155); +if (lean_is_exclusive(x_154)) { + lean_ctor_release(x_154, 0); + lean_ctor_release(x_154, 1); + x_156 = x_154; } else { - lean_dec_ref(x_117); - x_119 = lean_box(0); + lean_dec_ref(x_154); + x_156 = lean_box(0); } -x_120 = lean_box(0); -if (lean_is_scalar(x_119)) { - x_121 = lean_alloc_ctor(1, 2, 0); +x_157 = lean_box(0); +if (lean_is_scalar(x_156)) { + x_158 = lean_alloc_ctor(1, 2, 0); } else { - x_121 = x_119; - lean_ctor_set_tag(x_121, 1); + x_158 = x_156; + lean_ctor_set_tag(x_158, 1); } -lean_ctor_set(x_121, 0, x_115); -lean_ctor_set(x_121, 1, x_120); -x_122 = lean_st_ref_set(x_10, x_121, x_118); -x_123 = lean_ctor_get(x_122, 1); -lean_inc(x_123); -lean_dec(x_122); -x_124 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_123); -if (lean_obj_tag(x_124) == 0) +lean_ctor_set(x_158, 0, x_152); +lean_ctor_set(x_158, 1, x_157); +x_159 = lean_st_ref_set(x_10, x_158, x_155); +x_160 = lean_ctor_get(x_159, 1); +lean_inc(x_160); +lean_dec(x_159); +x_161 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices(x_144, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_160); +if (lean_obj_tag(x_161) == 0) { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_125 = lean_ctor_get(x_124, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_126 = x_124; +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_162 = lean_ctor_get(x_161, 1); +lean_inc(x_162); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_163 = x_161; } else { - lean_dec_ref(x_124); - x_126 = lean_box(0); + lean_dec_ref(x_161); + x_163 = lean_box(0); } -if (lean_is_scalar(x_116)) { - x_127 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_153)) { + x_164 = lean_alloc_ctor(1, 1, 0); } else { - x_127 = x_116; + x_164 = x_153; } -lean_ctor_set(x_127, 0, x_5); -if (lean_is_scalar(x_126)) { - x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_164, 0, x_5); +if (lean_is_scalar(x_163)) { + x_165 = lean_alloc_ctor(0, 2, 0); } else { - x_128 = x_126; + x_165 = x_163; } -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_125); -return x_128; +lean_ctor_set(x_165, 0, x_164); +lean_ctor_set(x_165, 1, x_162); +return x_165; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -lean_dec(x_116); +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +lean_dec(x_153); lean_dec(x_5); -x_129 = lean_ctor_get(x_124, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_124, 1); -lean_inc(x_130); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_131 = x_124; +x_166 = lean_ctor_get(x_161, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_161, 1); +lean_inc(x_167); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_168 = x_161; } else { - lean_dec_ref(x_124); - x_131 = lean_box(0); + lean_dec_ref(x_161); + x_168 = lean_box(0); } -if (lean_is_scalar(x_131)) { - x_132 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_168)) { + x_169 = lean_alloc_ctor(1, 2, 0); } else { - x_132 = x_131; + x_169 = x_168; } -lean_ctor_set(x_132, 0, x_129); -lean_ctor_set(x_132, 1, x_130); -return x_132; +lean_ctor_set(x_169, 0, x_166); +lean_ctor_set(x_169, 1, x_167); +return x_169; } } } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +lean_dec(x_144); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -12102,35 +16838,34 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); lean_dec(x_5); -x_133 = lean_ctor_get(x_108, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_108, 1); -lean_inc(x_134); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - lean_ctor_release(x_108, 1); - x_135 = x_108; +x_170 = lean_ctor_get(x_145, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_145, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_172 = x_145; } else { - lean_dec_ref(x_108); - x_135 = lean_box(0); + lean_dec_ref(x_145); + x_172 = lean_box(0); } -if (lean_is_scalar(x_135)) { - x_136 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_172)) { + x_173 = lean_alloc_ctor(1, 2, 0); } else { - x_136 = x_135; + x_173 = x_172; } -lean_ctor_set(x_136, 0, x_133); -lean_ctor_set(x_136, 1, x_134); -return x_136; +lean_ctor_set(x_173, 0, x_170); +lean_ctor_set(x_173, 1, x_171); +return x_173; } } } } else { -uint8_t x_155; +uint8_t x_191; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -12145,23 +16880,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_155 = !lean_is_exclusive(x_20); -if (x_155 == 0) +x_191 = !lean_is_exclusive(x_20); +if (x_191 == 0) { return x_20; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_20, 0); -x_157 = lean_ctor_get(x_20, 1); -lean_inc(x_157); -lean_inc(x_156); +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_ctor_get(x_20, 0); +x_193 = lean_ctor_get(x_20, 1); +lean_inc(x_193); +lean_inc(x_192); lean_dec(x_20); -x_158 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -return x_158; +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +return x_194; } } } @@ -12422,7 +17157,7 @@ lean_inc(x_24); lean_dec(x_23); x_25 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; x_26 = lean_mk_array(x_24, x_25); -x_27 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_27 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); x_28 = lean_st_ref_get(x_5, x_17); x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); @@ -12433,7 +17168,7 @@ x_31 = lean_ctor_get(x_29, 6); lean_inc(x_31); lean_dec(x_29); x_32 = lean_box(0); -x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; lean_inc(x_22); x_34 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(x_1, x_2, x_22, x_26, x_27, x_31, x_32, x_33, x_22, x_22, x_33, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_30); lean_dec(x_31); @@ -12577,7 +17312,7 @@ lean_inc(x_62); lean_dec(x_61); x_63 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_unassigned; x_64 = lean_mk_array(x_62, x_63); -x_65 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_65 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); x_66 = lean_st_ref_get(x_5, x_54); x_67 = lean_ctor_get(x_66, 0); lean_inc(x_67); @@ -12588,7 +17323,7 @@ x_69 = lean_ctor_get(x_67, 6); lean_inc(x_69); lean_dec(x_67); x_70 = lean_box(0); -x_71 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1; +x_71 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1; lean_inc(x_60); x_72 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main___spec__1(x_1, x_2, x_60, x_64, x_65, x_69, x_70, x_71, x_60, x_60, x_71, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_68); lean_dec(x_69); @@ -12767,7 +17502,7 @@ if (x_4 == 0) lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_1, 0); x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_alloc_ctor(1, 1, 0); +x_7 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_7, 0, x_5); lean_ctor_set(x_1, 1, x_2); lean_ctor_set(x_1, 0, x_7); @@ -12787,7 +17522,7 @@ x_10 = lean_ctor_get(x_1, 1); lean_inc(x_10); lean_inc(x_9); lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 1, 0); +x_11 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_11, 0, x_9); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -12855,7 +17590,7 @@ lean_object* x_24; lean_object* x_25; x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec(x_23); -x_25 = lean_alloc_ctor(1, 1, 0); +x_25 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_25, 0, x_17); lean_ctor_set(x_1, 1, x_2); lean_ctor_set(x_1, 0, x_25); @@ -12938,7 +17673,7 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_37, 1); lean_inc(x_38); lean_dec(x_37); -x_39 = lean_alloc_ctor(1, 1, 0); +x_39 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_39, 0, x_31); x_40 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_40, 0, x_39); @@ -13003,7 +17738,7 @@ static lean_object* _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___c lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__1; x_2 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(242u); +x_3 = lean_unsigned_to_nat(315u); x_4 = lean_unsigned_to_nat(22u); x_5 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -13020,7 +17755,7 @@ x_15 = !lean_is_exclusive(x_3); if (x_15 == 0) { uint8_t x_16; lean_object* x_17; -x_16 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +x_16 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); x_17 = lean_ctor_get(x_3, 0); lean_dec(x_17); lean_ctor_set(x_3, 0, x_1); @@ -13072,54 +17807,57 @@ return x_29; } else { -uint8_t x_30; lean_object* x_31; -x_30 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +uint8_t x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get_uint8(x_3, sizeof(void*)*2); +x_31 = lean_ctor_get(x_3, 1); +lean_inc(x_31); lean_dec(x_3); -x_31 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_31, 0, x_1); -lean_ctor_set_uint8(x_31, sizeof(void*)*1, x_30); +x_32 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set_uint8(x_32, sizeof(void*)*2, x_30); if (lean_obj_tag(x_14) == 0) { -lean_object* x_32; lean_object* x_33; -x_32 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; -x_33 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_32, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_33; +lean_object* x_33; lean_object* x_34; +x_33 = l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2; +x_34 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__2(x_33, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_34; } else { -lean_object* x_34; -x_34 = lean_ctor_get(x_14, 1); -lean_inc(x_34); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_14, 0); +lean_object* x_35; +x_35 = lean_ctor_get(x_14, 1); lean_inc(x_35); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_14, 0); +lean_inc(x_36); lean_dec(x_14); -x_36 = lean_box(0); -x_37 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_35, x_36, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_37; +x_37 = lean_box(0); +x_38 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_36, x_37, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_38; } else { if (x_30 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_14, 0); -lean_inc(x_38); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_14, 0); +lean_inc(x_39); lean_dec(x_14); -x_39 = lean_box(0); -x_40 = l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(x_34, x_39); -x_41 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_38, x_40, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_41; +x_40 = lean_box(0); +x_41 = l_List_mapTR_loop___at_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll___spec__1(x_35, x_40); +x_42 = l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_main(x_39, x_41, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_42; } else { -lean_object* x_42; lean_object* x_43; -lean_dec(x_34); -x_42 = lean_box(0); -x_43 = l_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll(x_14, x_42, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_43; +lean_object* x_43; lean_object* x_44; +lean_dec(x_35); +x_43 = lean_box(0); +x_44 = l_Lean_Meta_Grind_EMatch_ematchTheorem_tryAll(x_14, x_43, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_44; } } } @@ -13757,22 +18495,25 @@ if (x_13 == 0) { uint8_t x_14; lean_object* x_15; x_14 = 1; -lean_ctor_set_uint8(x_2, sizeof(void*)*1, x_14); +lean_ctor_set_uint8(x_2, sizeof(void*)*2, x_14); x_15 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_15; } else { -lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; x_16 = lean_ctor_get(x_2, 0); +x_17 = lean_ctor_get(x_2, 1); +lean_inc(x_17); lean_inc(x_16); lean_dec(x_2); -x_17 = 1; -x_18 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set_uint8(x_18, sizeof(void*)*1, x_17); -x_19 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; +x_18 = 1; +x_19 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_17); +lean_ctor_set_uint8(x_19, sizeof(void*)*2, x_18); +x_20 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_20; } } } @@ -13785,22 +18526,25 @@ if (x_14 == 0) { uint8_t x_15; lean_object* x_16; x_15 = 0; -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_15); +lean_ctor_set_uint8(x_3, sizeof(void*)*2, x_15); x_16 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_16; } else { -lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; x_17 = lean_ctor_get(x_3, 0); +x_18 = lean_ctor_get(x_3, 1); +lean_inc(x_18); lean_inc(x_17); lean_dec(x_3); -x_18 = 0; -x_19 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set_uint8(x_19, sizeof(void*)*1, x_18); -x_20 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_20; +x_19 = 0; +x_20 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_20, 1, x_18); +lean_ctor_set_uint8(x_20, sizeof(void*)*2, x_19); +x_21 = l_Lean_PersistentArray_forM___at_Lean_Meta_Grind_EMatch_ematchTheorems___spec__1(x_1, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_21; } } } @@ -13843,74 +18587,107 @@ return x_5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematch(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_unbox(x_11); -lean_dec(x_11); -if (x_12 == 0) +uint8_t x_10; lean_object* x_11; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_unbox(x_84); +if (x_85 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_13 = lean_ctor_get(x_10, 1); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +lean_dec(x_84); +x_86 = lean_ctor_get(x_83, 1); +lean_inc(x_86); +lean_dec(x_83); +x_87 = l_Lean_Meta_Grind_checkMaxEmatchExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_86); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = lean_unbox(x_88); +lean_dec(x_88); +x_10 = x_90; +x_11 = x_89; +goto block_82; +} +else +{ +lean_object* x_91; uint8_t x_92; +x_91 = lean_ctor_get(x_83, 1); +lean_inc(x_91); +lean_dec(x_83); +x_92 = lean_unbox(x_84); +lean_dec(x_84); +x_10 = x_92; +x_11 = x_91; +goto block_82; +} +block_82: +{ +if (x_10 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_12 = lean_st_ref_get(x_1, x_11); +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_st_ref_get(x_1, x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_st_ref_get(x_1, x_14); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_st_ref_get(x_1, x_16); -x_18 = lean_ctor_get(x_17, 0); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_13, 8); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_13); +x_19 = lean_ctor_get(x_16, 9); lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_15, 8); -lean_inc(x_20); -lean_dec(x_15); -x_21 = lean_ctor_get(x_18, 9); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__1___boxed), 12, 1); +lean_dec(x_16); +x_20 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__1___boxed), 12, 1); +lean_closure_set(x_20, 0, x_18); +x_21 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__2___boxed), 13, 1); +lean_closure_set(x_21, 0, x_19); +x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); lean_closure_set(x_22, 0, x_20); -x_23 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematch___lambda__2___boxed), 13, 1); -lean_closure_set(x_23, 0, x_21); -x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__11___rarg), 13, 2); -lean_closure_set(x_24, 0, x_22); -lean_closure_set(x_24, 1, x_23); +lean_closure_set(x_22, 1, x_21); lean_inc(x_1); -x_25 = l_Lean_Meta_Grind_EMatch_M_run_x27___rarg(x_24, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -if (lean_obj_tag(x_25) == 0) +x_23 = l_Lean_Meta_Grind_EMatch_M_run_x27___rarg(x_22, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_26 = lean_ctor_get(x_25, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_st_ref_take(x_1, x_24); +x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); lean_dec(x_25); -x_27 = lean_st_ref_take(x_1, x_26); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = !lean_is_exclusive(x_28); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_31 = lean_ctor_get(x_28, 6); -x_32 = lean_ctor_get(x_28, 8); -x_33 = lean_ctor_get(x_28, 9); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_add(x_31, x_34); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_29 = lean_ctor_get(x_26, 6); +x_30 = lean_ctor_get(x_26, 8); +x_31 = lean_ctor_get(x_26, 9); +x_32 = lean_ctor_get(x_26, 12); +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_add(x_29, x_33); +lean_dec(x_29); +x_35 = l_Lean_PersistentArray_append___rarg(x_30, x_31); lean_dec(x_31); -x_36 = l_Lean_PersistentArray_append___rarg(x_32, x_33); -lean_dec(x_33); +x_36 = lean_nat_add(x_32, x_33); +lean_dec(x_32); x_37 = l_Lean_Meta_Grind_ematch___closed__3; -lean_ctor_set(x_28, 9, x_37); -lean_ctor_set(x_28, 8, x_36); -lean_ctor_set(x_28, 6, x_35); -x_38 = lean_st_ref_set(x_1, x_28, x_29); +lean_ctor_set(x_26, 12, x_36); +lean_ctor_set(x_26, 9, x_37); +lean_ctor_set(x_26, 8, x_35); +lean_ctor_set(x_26, 6, x_34); +x_38 = lean_st_ref_set(x_1, x_26, x_27); lean_dec(x_1); x_39 = !lean_is_exclusive(x_38); if (x_39 == 0) @@ -13937,22 +18714,32 @@ return x_44; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_45 = lean_ctor_get(x_28, 0); -x_46 = lean_ctor_get(x_28, 1); -x_47 = lean_ctor_get(x_28, 2); -x_48 = lean_ctor_get(x_28, 3); -x_49 = lean_ctor_get(x_28, 4); -x_50 = lean_ctor_get(x_28, 5); -x_51 = lean_ctor_get_uint8(x_28, sizeof(void*)*14); -x_52 = lean_ctor_get(x_28, 6); -x_53 = lean_ctor_get(x_28, 7); -x_54 = lean_ctor_get(x_28, 8); -x_55 = lean_ctor_get(x_28, 9); -x_56 = lean_ctor_get(x_28, 10); -x_57 = lean_ctor_get(x_28, 11); -x_58 = lean_ctor_get(x_28, 12); -x_59 = lean_ctor_get(x_28, 13); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_45 = lean_ctor_get(x_26, 0); +x_46 = lean_ctor_get(x_26, 1); +x_47 = lean_ctor_get(x_26, 2); +x_48 = lean_ctor_get(x_26, 3); +x_49 = lean_ctor_get(x_26, 4); +x_50 = lean_ctor_get(x_26, 5); +x_51 = lean_ctor_get_uint8(x_26, sizeof(void*)*19); +x_52 = lean_ctor_get(x_26, 6); +x_53 = lean_ctor_get(x_26, 7); +x_54 = lean_ctor_get(x_26, 8); +x_55 = lean_ctor_get(x_26, 9); +x_56 = lean_ctor_get(x_26, 10); +x_57 = lean_ctor_get(x_26, 11); +x_58 = lean_ctor_get(x_26, 12); +x_59 = lean_ctor_get(x_26, 13); +x_60 = lean_ctor_get(x_26, 14); +x_61 = lean_ctor_get(x_26, 15); +x_62 = lean_ctor_get(x_26, 16); +x_63 = lean_ctor_get(x_26, 17); +x_64 = lean_ctor_get(x_26, 18); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); lean_inc(x_59); lean_inc(x_58); lean_inc(x_57); @@ -13967,79 +18754,86 @@ lean_inc(x_48); lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); -lean_dec(x_28); -x_60 = lean_unsigned_to_nat(1u); -x_61 = lean_nat_add(x_52, x_60); +lean_dec(x_26); +x_65 = lean_unsigned_to_nat(1u); +x_66 = lean_nat_add(x_52, x_65); lean_dec(x_52); -x_62 = l_Lean_PersistentArray_append___rarg(x_54, x_55); +x_67 = l_Lean_PersistentArray_append___rarg(x_54, x_55); lean_dec(x_55); -x_63 = l_Lean_Meta_Grind_ematch___closed__3; -x_64 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_64, 0, x_45); -lean_ctor_set(x_64, 1, x_46); -lean_ctor_set(x_64, 2, x_47); -lean_ctor_set(x_64, 3, x_48); -lean_ctor_set(x_64, 4, x_49); -lean_ctor_set(x_64, 5, x_50); -lean_ctor_set(x_64, 6, x_61); -lean_ctor_set(x_64, 7, x_53); -lean_ctor_set(x_64, 8, x_62); -lean_ctor_set(x_64, 9, x_63); -lean_ctor_set(x_64, 10, x_56); -lean_ctor_set(x_64, 11, x_57); -lean_ctor_set(x_64, 12, x_58); -lean_ctor_set(x_64, 13, x_59); -lean_ctor_set_uint8(x_64, sizeof(void*)*14, x_51); -x_65 = lean_st_ref_set(x_1, x_64, x_29); -lean_dec(x_1); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_67 = x_65; +x_68 = lean_nat_add(x_58, x_65); +lean_dec(x_58); +x_69 = l_Lean_Meta_Grind_ematch___closed__3; +x_70 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_70, 0, x_45); +lean_ctor_set(x_70, 1, x_46); +lean_ctor_set(x_70, 2, x_47); +lean_ctor_set(x_70, 3, x_48); +lean_ctor_set(x_70, 4, x_49); +lean_ctor_set(x_70, 5, x_50); +lean_ctor_set(x_70, 6, x_66); +lean_ctor_set(x_70, 7, x_53); +lean_ctor_set(x_70, 8, x_67); +lean_ctor_set(x_70, 9, x_69); +lean_ctor_set(x_70, 10, x_56); +lean_ctor_set(x_70, 11, x_57); +lean_ctor_set(x_70, 12, x_68); +lean_ctor_set(x_70, 13, x_59); +lean_ctor_set(x_70, 14, x_60); +lean_ctor_set(x_70, 15, x_61); +lean_ctor_set(x_70, 16, x_62); +lean_ctor_set(x_70, 17, x_63); +lean_ctor_set(x_70, 18, x_64); +lean_ctor_set_uint8(x_70, sizeof(void*)*19, x_51); +x_71 = lean_st_ref_set(x_1, x_70, x_27); +lean_dec(x_1); +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_73 = x_71; } else { - lean_dec_ref(x_65); - x_67 = lean_box(0); + lean_dec_ref(x_71); + x_73 = lean_box(0); } -x_68 = lean_box(0); -if (lean_is_scalar(x_67)) { - x_69 = lean_alloc_ctor(0, 2, 0); +x_74 = lean_box(0); +if (lean_is_scalar(x_73)) { + x_75 = lean_alloc_ctor(0, 2, 0); } else { - x_69 = x_67; + x_75 = x_73; } -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_66); -return x_69; +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_72); +return x_75; } } else { -uint8_t x_70; +uint8_t x_76; lean_dec(x_1); -x_70 = !lean_is_exclusive(x_25); -if (x_70 == 0) +x_76 = !lean_is_exclusive(x_23); +if (x_76 == 0) { -return x_25; +return x_23; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_25, 0); -x_72 = lean_ctor_get(x_25, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_25); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -return x_73; +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_23, 0); +x_78 = lean_ctor_get(x_23, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_23); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; } } } else { -uint8_t x_74; +lean_object* x_80; lean_object* x_81; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -14048,27 +18842,11 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_74 = !lean_is_exclusive(x_10); -if (x_74 == 0) -{ -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_10, 0); -lean_dec(x_75); -x_76 = lean_box(0); -lean_ctor_set(x_10, 0, x_76); -return x_10; -} -else -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_10, 1); -lean_inc(x_77); -lean_dec(x_10); -x_78 = lean_box(0); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_77); -return x_79; +x_80 = lean_box(0); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_11); +return x_81; } } } @@ -14092,7 +18870,7 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; uint8_t x_11; @@ -14117,7 +18895,7 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; @@ -14221,7 +18999,7 @@ return x_32; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; @@ -14247,90 +19025,40 @@ return x_13; } } } -static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext_x3f), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1; -x_12 = l_Lean_Meta_Grind_iterate(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_12, 0, x_15); -return x_12; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_12, 0); -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_12); -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_16); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -} -else -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_12); -if (x_20 == 0) -{ +x_11 = l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1; +x_12 = l_Lean_Meta_Grind_GrindTactic_iterate(x_11, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 0); -x_22 = lean_ctor_get(x_12, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_12); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} } -static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__2), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert___lambda__2), 9, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_ematchAndAssert___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert___lambda__3___boxed), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -14338,13 +19066,13 @@ x_10 = lean_ctor_get(x_1, 11); lean_inc(x_10); x_11 = lean_ctor_get(x_1, 0); lean_inc(x_11); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1___boxed), 9, 1); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert___lambda__1___boxed), 9, 1); lean_closure_set(x_12, 0, x_1); -x_13 = l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1; +x_13 = l_Lean_Meta_Grind_ematchAndAssert___closed__1; x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); lean_closure_set(x_14, 0, x_12); lean_closure_set(x_14, 1, x_13); -x_15 = l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2; +x_15 = l_Lean_Meta_Grind_ematchAndAssert___closed__2; x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); lean_closure_set(x_16, 0, x_14); lean_closure_set(x_16, 1, x_15); @@ -14374,8 +19102,8 @@ if (x_22 == 0) { lean_object* x_23; lean_object* x_24; lean_free_object(x_17); -x_23 = lean_box(0); -x_24 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(x_19, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext), 9, 0); +x_24 = l_Lean_Meta_Grind_GrindTactic_iterate(x_23, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_20); return x_24; } else @@ -14410,8 +19138,8 @@ lean_dec(x_28); if (x_29 == 0) { lean_object* x_30; lean_object* x_31; -x_30 = lean_box(0); -x_31 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(x_26, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); +x_30 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext), 9, 0); +x_31 = l_Lean_Meta_Grind_GrindTactic_iterate(x_30, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); return x_31; } else @@ -14465,11 +19193,11 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_ematchAndAssert___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -14480,11 +19208,11 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_ematchAndAssert___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -14495,11 +19223,11 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchAndAssert___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_Lean_Meta_Grind_ematchAndAssert___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_2); return x_11; } @@ -14508,7 +19236,7 @@ static lean_object* _init_l_Lean_Meta_Grind_ematchStar___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert_x3f), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchAndAssert), 9, 0); return x_1; } } @@ -14517,12 +19245,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ematchStar(lean_object* x_1, lean_obj { lean_object* x_10; lean_object* x_11; x_10 = l_Lean_Meta_Grind_ematchStar___closed__1; -x_11 = l_Lean_Meta_Grind_iterate(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_Grind_GrindTactic_iterate(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_DoNotSimp(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(uint8_t builtin, lean_object* w) { lean_object * res; @@ -14534,6 +19263,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Intro(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_DoNotSimp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__1); l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2 = _init_l_Lean_Meta_Grind_EMatch_instInhabitedCnstr___closed__2(); @@ -14594,8 +19326,36 @@ l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f__ lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__3); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_assign_x3f___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_matchArg_x3f___closed__8); l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1(); lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processContinue___spec__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_annotateMatchEqnType___closed__2); l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1(); lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__1); l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__2(); @@ -14604,12 +19364,6 @@ l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Gr lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__3); l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4 = _init_l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4(); lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__5___closed__4); -l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__1); -l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__2); -l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3 = _init_l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__3___closed__3); l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__1(); l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2(); lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_addNewInstance___spec__7___closed__2); @@ -14657,48 +19411,46 @@ l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_ lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__8); l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9(); lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__9); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__10); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__3___closed__11); l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__1); l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___spec__5___closed__2); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__1(); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__2); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__3); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__4); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__3___closed__5); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__5___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__1(); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__2___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__6___closed__7); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__1); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__2); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__3); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__4); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__5); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__6); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__7); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__7___closed__8); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__1); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___lambda__8___closed__2); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__1); l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_instantiateTheorem___closed__2(); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___lambda__2___closed__1); -l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatch_0__Lean_Meta_Grind_EMatch_processChoices___closed__1); l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__1); l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_EMatch_ematchTheorem___lambda__1___closed__2(); @@ -14709,12 +19461,12 @@ l_Lean_Meta_Grind_ematch___closed__2 = _init_l_Lean_Meta_Grind_ematch___closed__ lean_mark_persistent(l_Lean_Meta_Grind_ematch___closed__2); l_Lean_Meta_Grind_ematch___closed__3 = _init_l_Lean_Meta_Grind_ematch___closed__3(); lean_mark_persistent(l_Lean_Meta_Grind_ematch___closed__3); -l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___lambda__4___closed__1); -l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__1); -l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2 = _init_l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert_x3f___closed__2); +l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert___lambda__4___closed__1); +l_Lean_Meta_Grind_ematchAndAssert___closed__1 = _init_l_Lean_Meta_Grind_ematchAndAssert___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert___closed__1); +l_Lean_Meta_Grind_ematchAndAssert___closed__2 = _init_l_Lean_Meta_Grind_ematchAndAssert___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_ematchAndAssert___closed__2); l_Lean_Meta_Grind_ematchStar___closed__1 = _init_l_Lean_Meta_Grind_ematchStar___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_ematchStar___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c index 3186aa6411c1..37bcdd8f396e 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/EMatchTheorem.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.EMatchTheorem -// Imports: Lean.HeadIndex Lean.PrettyPrinter Lean.Util.FoldConsts Lean.Util.CollectFVars Lean.Meta.Basic Lean.Meta.InferType +// Imports: Init.Grind.Util Init.Grind.Tactics Lean.HeadIndex Lean.PrettyPrinter Lean.Util.FoldConsts Lean.Util.CollectFVars Lean.Meta.Basic Lean.Meta.InferType Lean.Meta.Eqns Lean.Meta.Tactic.Grind.Util #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,185 +13,275 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt; lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21; lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__4; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15; +static lean_object* l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableOrigin___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3; static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_mkGroundPattern___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_object*); static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_contains___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_List_all___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; -static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; +static lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Array_contains___at_Lean_Meta_arrowDomainsN___spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); lean_object* l_Lean_indentD(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isPatternDontCare(lean_object*); +uint8_t l_Lean_Exception_isInterrupt(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2; uint64_t lean_uint64_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); +lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_isOffset_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1; +lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedOrigin; size_t lean_uint64_to_usize(uint64_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1; lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__5; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__6; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; size_t lean_usize_mul(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3; lean_object* l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__3; lean_object* lean_mk_array(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279_(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqOrigin__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__8; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16; static lean_object* l_Lean_Meta_Grind_mkGroundPattern___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; lean_object* l_Lean_Expr_bvar___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__11; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Meta_registerCoercion___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorems; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666_(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2(uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instReprOrigin; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instReprOrigin___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__8; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1(lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_EMatchTheorems_isErased(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isPatternDontCare___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5; uint8_t l_Lean_Expr_isBVar(lean_object*); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__2; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_mkGroundPattern___closed__2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key___boxed(lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__3; static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3; +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__2; size_t lean_usize_of_nat(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_expr_abstract(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14; static lean_object* l_Lean_Meta_Grind_ppPattern___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_ConstantInfo_isTheorem(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2___boxed(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); @@ -199,1323 +289,1354 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lea lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7; lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__5; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10; +static lean_object* l_Lean_Meta_Grind_preprocessPattern___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; -LEAN_EXPORT uint8_t l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getKind(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1; +static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(lean_object*, lean_object*); +lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2___boxed(lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems(lean_object*); +lean_object* l_Lean_Meta_getEqnsFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8; static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1; +lean_object* l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16; +static lean_object* l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); +lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540____boxed(lean_object*, lean_object*); +lean_object* l_Lean_ScopedEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; uint8_t l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqOrigin; static lean_object* l_Lean_Meta_Grind_ppPattern___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6; lean_object* l_Lean_Expr_toHeadIndex(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; +lean_object* l_panic___at_String_fromUTF8_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__11; +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__4; lean_object* l_Lean_mkConstWithLevelParams___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9; lean_object* l_Lean_Expr_constName_x21(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1; lean_object* l_Lean_MessageData_ofConst(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9; lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_grind_normalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6; lean_object* l_Lean_mkAnnotation(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_preprocessPattern___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____boxed(lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Expr_instantiateLevelParamsArray(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isArrow(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqTheoremKind; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure(uint8_t); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1; -static lean_object* l_Lean_Meta_Grind_Origin_key___closed__2; +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1; lean_object* l_Lean_MessageData_ofExpr(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1; +static lean_object* l_Lean_Meta_Grind_mkOffsetPattern___closed__3; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(uint8_t, uint8_t); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_isErased___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12; +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20; +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableOrigin(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4; +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__9; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; -lean_object* l_Lean_PersistentHashMap_instInhabited(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_normalizePattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4; +static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_EMatchTheorems_contains(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2; +lean_object* l_Lean_mkRawNatLit(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_id___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion(lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; uint64_t l_Lean_Name_hash___override(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5; uint64_t lean_uint64_xor(uint64_t, uint64_t); extern lean_object* l_Id_instMonad; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1; lean_object* l_Lean_annotation_x3f(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instInhabitedTheoremKind; +static uint64_t l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Meta_isTypeFormer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2; +lean_object* l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11; lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(lean_object*, size_t, lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; static lean_object* l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1; +lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_key___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2; +uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__10; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4; +static lean_object* l_Lean_Meta_Grind_instBEqOrigin___closed__1; +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2; +static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Expr_fvar___override(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; size_t lean_array_size(lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__12; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f(lean_object*, lean_object*); lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___boxed(lean_object*); uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_preprocessPattern___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_instBEqTheoremKind___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__10; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqOrigin__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17; -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12; -extern lean_object* l_Lean_Name_instBEq; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_erase(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(uint8_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkOffsetPattern(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10; static lean_object* l_Lean_Meta_Grind_ppPattern___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); lean_object* l_Lean_MessageData_bracket(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1; uint8_t l_Lean_Expr_isConst(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5; +uint8_t l_Lean_Exception_isRuntime(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchEqTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_HeadIndex_hash(lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); static lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_ppPattern___closed__6; lean_object* l_Lean_Meta_mkConstWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___closed__1; +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5; +static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +uint64_t l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(lean_object*); +lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_MessageData_ofName(lean_object*); +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__4; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1; -static lean_object* l_Lean_Meta_Grind_addEMatchTheorem___closed__7; +static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9; lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1; -extern lean_object* l_Lean_instHashableName; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1; lean_object* l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3; lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__2() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.decl", 27, 27); +x_1 = lean_mk_string_unchecked("offset", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Meta_Grind_mkOffsetPattern___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__4; +x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkOffsetPattern(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(2u); -x_2 = lean_nat_to_int(x_1); -return x_2; +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Lean_mkRawNatLit(x_2); +x_4 = l_Lean_Meta_Grind_mkOffsetPattern___closed__5; +x_5 = l_Lean_mkAppB(x_4, x_1, x_3); +return x_5; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_nat_to_int(x_1); -return x_2; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = l_Lean_Meta_Grind_mkOffsetPattern(x_1, x_2); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_11, 0, x_10); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.fvar", 27, 27); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7() { -_start: +uint8_t x_8; +x_8 = lean_expr_eqv(x_2, x_1); +if (x_8 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +switch (lean_obj_tag(x_2)) { +case 6: +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_9 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_9, 0, x_2); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +return x_10; } +case 7: +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_2); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_7); +return x_12; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8() { -_start: +case 8: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_13; lean_object* x_14; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_2); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_7); +return x_14; } +default: +{ +lean_object* x_15; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_15 = l_Lean_Meta_isOffset_x3f(x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_2); +x_20 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_15, 0, x_20); +return x_15; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.stx", 26, 26); -return x_1; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_2); +x_23 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +return x_24; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +uint8_t x_25; +lean_dec(x_2); +x_25 = !lean_is_exclusive(x_16); +if (x_25 == 0) +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_15); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_27 = lean_ctor_get(x_16, 0); +x_28 = lean_ctor_get(x_15, 1); +x_29 = lean_ctor_get(x_15, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_27, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_unsigned_to_nat(0u); +x_33 = lean_nat_dec_eq(x_31, x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +lean_free_object(x_15); +lean_free_object(x_16); +x_34 = lean_box(0); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(x_30, x_31, x_34, x_3, x_4, x_5, x_6, x_28); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_35; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10; -x_2 = lean_box(1); -x_3 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_36; +lean_dec(x_31); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_ctor_set(x_16, 0, x_30); +x_36 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_36, 0, x_16); +lean_ctor_set(x_15, 0, x_36); +return x_15; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.other", 28, 28); -return x_1; -} +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_37 = lean_ctor_get(x_16, 0); +x_38 = lean_ctor_get(x_15, 1); +lean_inc(x_38); +lean_dec(x_15); +x_39 = lean_ctor_get(x_37, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = lean_unsigned_to_nat(0u); +x_42 = lean_nat_dec_eq(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +lean_free_object(x_16); +x_43 = lean_box(0); +x_44 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(x_39, x_40, x_43, x_3, x_4, x_5, x_6, x_38); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_44; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_45; lean_object* x_46; +lean_dec(x_40); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_ctor_set(x_16, 0, x_39); +x_45 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_45, 0, x_16); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_38); +return x_46; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; -x_3 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_47 = lean_ctor_get(x_16, 0); +lean_inc(x_47); +lean_dec(x_16); +x_48 = lean_ctor_get(x_15, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + lean_ctor_release(x_15, 1); + x_49 = x_15; +} else { + lean_dec_ref(x_15); + x_49 = lean_box(0); } +x_50 = lean_ctor_get(x_47, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_47, 1); +lean_inc(x_51); +lean_dec(x_47); +x_52 = lean_unsigned_to_nat(0u); +x_53 = lean_nat_dec_eq(x_51, x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +lean_dec(x_49); +x_54 = lean_box(0); +x_55 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(x_50, x_51, x_54, x_3, x_4, x_5, x_6, x_48); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_55; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15() { -_start: +else { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14; -x_2 = 0; -x_3 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); -return x_3; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_51); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_50); +x_57 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_57, 0, x_56); +if (lean_is_scalar(x_49)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_49; } +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_48); +return x_58; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13; -x_3 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16; -x_2 = 0; -x_3 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(lean_object* x_1, lean_object* x_2) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -lean_dec(x_1); -x_4 = lean_unsigned_to_nat(1024u); -x_5 = lean_nat_dec_le(x_4, x_2); -x_6 = l_Lean_Name_reprPrec(x_3, x_4); -x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3; -x_8 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -if (x_5 == 0) -{ -lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; -x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_10 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -x_11 = 0; -x_12 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); -x_13 = l_Repr_addAppParen(x_12, x_2); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -x_15 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_8); -x_16 = 0; -x_17 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_16); -x_18 = l_Repr_addAppParen(x_17, x_2); -return x_18; -} -} -case 1: -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_19 = lean_ctor_get(x_1, 0); -lean_inc(x_19); -lean_dec(x_1); -x_20 = lean_unsigned_to_nat(1024u); -x_21 = lean_nat_dec_le(x_20, x_2); -x_22 = l_Lean_Name_reprPrec(x_19, x_20); -x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8; -x_24 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -if (x_21 == 0) -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_25 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_26 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -x_27 = 0; -x_28 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); -x_29 = l_Repr_addAppParen(x_28, x_2); -return x_29; -} -else -{ -lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; -x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -x_31 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_24); -x_32 = 0; -x_33 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); -x_34 = l_Repr_addAppParen(x_33, x_2); -return x_34; -} -} -case 2: -{ -uint8_t x_35; -x_35 = !lean_is_exclusive(x_1); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_36 = lean_ctor_get(x_1, 0); -x_37 = lean_ctor_get(x_1, 1); -x_38 = lean_unsigned_to_nat(1024u); -x_39 = lean_nat_dec_le(x_38, x_2); -x_40 = l_Lean_Name_reprPrec(x_36, x_38); -x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; -lean_ctor_set_tag(x_1, 5); -lean_ctor_set(x_1, 1, x_40); -lean_ctor_set(x_1, 0, x_41); -x_42 = lean_box(1); -x_43 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_42); -x_44 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_37, x_38); -x_45 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -if (x_39 == 0) -{ -lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; -x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_47 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_45); -x_48 = 0; -x_49 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_48); -x_50 = l_Repr_addAppParen(x_49, x_2); -return x_50; -} -else -{ -lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; -x_51 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -x_52 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_45); -x_53 = 0; -x_54 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_53); -x_55 = l_Repr_addAppParen(x_54, x_2); -return x_55; -} } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_56 = lean_ctor_get(x_1, 0); -x_57 = lean_ctor_get(x_1, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_1); -x_58 = lean_unsigned_to_nat(1024u); -x_59 = lean_nat_dec_le(x_58, x_2); -x_60 = l_Lean_Name_reprPrec(x_56, x_58); -x_61 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11; -x_62 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_60); -x_63 = lean_box(1); -x_64 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -x_65 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_57, x_58); -x_66 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); +uint8_t x_59; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_59 = !lean_is_exclusive(x_15); if (x_59 == 0) { -lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; -x_67 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4; -x_68 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = 0; -x_70 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set_uint8(x_70, sizeof(void*)*1, x_69); -x_71 = l_Repr_addAppParen(x_70, x_2); -return x_71; +return x_15; } else { -lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; -x_72 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5; -x_73 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_66); -x_74 = 0; -x_75 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set_uint8(x_75, sizeof(void*)*1, x_74); -x_76 = l_Repr_addAppParen(x_75, x_2); -return x_76; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_15, 0); +x_61 = lean_ctor_get(x_15, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_15); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} } } } -default: -{ -lean_object* x_77; uint8_t x_78; -x_77 = lean_unsigned_to_nat(1024u); -x_78 = lean_nat_dec_le(x_77, x_2); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; -x_79 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15; -x_80 = l_Repr_addAppParen(x_79, x_2); -return x_80; } else { -lean_object* x_81; lean_object* x_82; -x_81 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17; -x_82 = l_Repr_addAppParen(x_81, x_2); -return x_82; -} -} +lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_63 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_63, 0, x_2); +x_64 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_64, 0, x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_7); +return x_65; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51_(x_1, x_2); -lean_dec(x_2); -return x_3; +lean_object* x_7; lean_object* x_8; +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } } -static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instReprOrigin___closed__1; -return x_1; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_inc(x_1); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2___boxed), 7, 1); +lean_closure_set(x_7, 0, x_1); +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1; +x_9 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_1, x_7, x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; } } -static lean_object* _init_l_Lean_Meta_Grind_Origin_key___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("other", 5, 5); -return x_1; +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; } } -static lean_object* _init_l_Lean_Meta_Grind_Origin_key___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_Origin_key___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_1); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -if (lean_obj_tag(x_1) == 3) -{ -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_Origin_key___closed__2; -return x_2; -} -else -{ -lean_object* x_3; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -return x_3; -} +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_Origin_key(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: +if (lean_obj_tag(x_2) == 9) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_1, 0); +lean_object* x_3; +x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -lean_dec(x_1); -x_4 = lean_ctor_get(x_3, 1); -lean_inc(x_4); -lean_dec(x_3); -x_5 = l_Lean_MessageData_ofConst(x_2); -x_6 = lean_apply_2(x_4, lean_box(0), x_5); -return x_6; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("[unknown]", 9, 9); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__2() { -_start: +lean_dec(x_2); +if (lean_obj_tag(x_3) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Origin_pp___rarg___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__3() { -_start: +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_Origin_pp___rarg___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; -} +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_5); +lean_ctor_set_tag(x_3, 1); +lean_ctor_set(x_3, 0, x_6); +return x_3; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -switch (lean_obj_tag(x_4)) { -case 0: +else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -lean_dec(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -lean_inc(x_1); -x_7 = l_Lean_mkConstWithLevelParams___rarg(x_1, x_2, x_3, x_5); -x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1), 2, 1); -lean_closure_set(x_8, 0, x_1); -x_9 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_8); +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec(x_3); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); return x_9; } -case 1: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_3); -lean_dec(x_2); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -lean_dec(x_1); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Expr_fvar___override(x_10); -x_14 = l_Lean_MessageData_ofExpr(x_13); -x_15 = lean_apply_2(x_12, lean_box(0), x_14); -return x_15; } -case 2: +else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_10; lean_dec(x_3); -lean_dec(x_2); -x_16 = lean_ctor_get(x_4, 1); -lean_inc(x_16); -lean_dec(x_4); -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); lean_dec(x_1); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_Lean_MessageData_ofSyntax(x_16); -x_20 = lean_apply_2(x_18, lean_box(0), x_19); -return x_20; +x_10 = lean_box(0); +return x_10; } -default: +} +else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_3); +lean_object* x_11; lean_dec(x_2); -x_21 = lean_ctor_get(x_1, 0); -lean_inc(x_21); lean_dec(x_1); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Meta_Grind_Origin_pp___rarg___closed__3; -x_24 = lean_apply_2(x_22, lean_box(0), x_23); -return x_24; -} +x_11 = lean_box(0); +return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg), 4, 0); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Expr_bvar___override(x_1); +x_2 = lean_box(0); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1; -x_3 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2; -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; -x_6 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_6, 0, x_2); -lean_ctor_set(x_6, 1, x_3); -lean_ctor_set(x_6, 2, x_4); -lean_ctor_set(x_6, 3, x_1); -lean_ctor_set(x_6, 4, x_1); -lean_ctor_set(x_6, 5, x_5); -return x_6; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__1), 2, 0); +return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem() { +static lean_object* _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Name_instBEq; -x_2 = l_Lean_instHashableName; -x_3 = l_Lean_PersistentHashMap_instInhabited(lean_box(0), lean_box(0), x_1, x_2); +x_1 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2; +x_2 = lean_box(0); +x_3 = lean_apply_1(x_1, x_2); return x_3; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Id_instMonad; -x_2 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(lean_object* x_1) { -_start: +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1; +x_3 = l_Lean_Expr_cleanupAnnotations(x_1); +x_4 = l_Lean_Expr_isApp(x_3); +if (x_4 == 0) { -lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} +lean_object* x_5; +lean_dec(x_3); +x_5 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3; +return x_5; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); -if (x_7 == 0) +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = l_Lean_Expr_appArg(x_3, lean_box(0)); +x_7 = l_Lean_Expr_appFnCleanup(x_3, lean_box(0)); +x_8 = l_Lean_Expr_isApp(x_7); +if (x_8 == 0) { -lean_object* x_8; -lean_dec(x_4); -x_8 = lean_box(0); -return x_8; +lean_object* x_9; +lean_dec(x_7); +lean_dec(x_6); +x_9 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3; +return x_9; } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_name_eq(x_5, x_9); -lean_dec(x_9); -if (x_10 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = l_Lean_Expr_appArg(x_7, lean_box(0)); +x_11 = l_Lean_Expr_appFnCleanup(x_7, lean_box(0)); +x_12 = l_Lean_Meta_Grind_mkOffsetPattern___closed__4; +x_13 = l_Lean_Expr_isConstOf(x_11, x_12); +lean_dec(x_11); +if (x_13 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_12; -goto _start; +lean_object* x_14; +lean_dec(x_10); +lean_dec(x_6); +x_14 = l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3; +return x_14; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_array_fget(x_2, x_4); -lean_dec(x_4); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); +lean_object* x_15; +x_15 = lean_apply_2(x_2, x_10, x_6); return x_15; } } } } -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1() { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2___boxed(lean_object* x_1) { _start: { -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = 5; -x_3 = lean_usize_shift_left(x_1, x_2); -return x_3; +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_isOffsetPattern_x3f___lambda__2(x_1); +lean_dec(x_1); +return x_2; } } -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1() { _start: { -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; -x_3 = lean_usize_sub(x_2, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__3___boxed), 6, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2() { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +lean_object* x_7; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_box(2); -x_11 = lean_array_get(x_10, x_5, x_9); -lean_dec(x_9); -lean_dec(x_5); -switch (lean_obj_tag(x_11)) { -case 0: +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1; +x_11 = l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2; +x_12 = 0; +x_13 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_8, x_10, x_11, x_12, x_12, x_2, x_3, x_4, x_5, x_9); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_name_eq(x_3, x_12); -lean_dec(x_12); +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { -lean_object* x_15; -lean_dec(x_13); -lean_free_object(x_1); -x_15 = lean_box(0); -return x_15; +return x_13; } else { -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_13); -return x_1; -} -} -case 1: -{ -lean_object* x_16; size_t x_17; -lean_free_object(x_1); -x_16 = lean_ctor_get(x_11, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_usize_shift_right(x_2, x_6); -x_1 = x_16; -x_2 = x_17; -goto _start; -} -default: -{ -lean_object* x_19; -lean_free_object(x_1); -x_19 = lean_box(0); -return x_19; -} +lean_inc(x_15); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } else { -lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); -lean_dec(x_1); -x_21 = 5; -x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; -x_23 = lean_usize_land(x_2, x_22); -x_24 = lean_usize_to_nat(x_23); -x_25 = lean_box(2); -x_26 = lean_array_get(x_25, x_20, x_24); -lean_dec(x_24); -lean_dec(x_20); -switch (lean_obj_tag(x_26)) { -case 0: -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_name_eq(x_3, x_27); -lean_dec(x_27); -if (x_29 == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) { -lean_object* x_30; -lean_dec(x_28); -x_30 = lean_box(0); -return x_30; +return x_13; } else { -lean_object* x_31; -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_28); -return x_31; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 0); +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_13); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -case 1: -{ -lean_object* x_32; size_t x_33; -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_usize_shift_right(x_2, x_21); -x_1 = x_32; -x_2 = x_33; -goto _start; } -default: +else { -lean_object* x_35; -x_35 = lean_box(0); -return x_35; -} -} -} +uint8_t x_22; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_22 = !lean_is_exclusive(x_7); +if (x_22 == 0) +{ +return x_7; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_1, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 1); -lean_inc(x_37); -lean_dec(x_1); -x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_36, x_37, lean_box(0), x_38, x_3); -lean_dec(x_37); -lean_dec(x_36); -return x_39; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_7, 0); +x_24 = lean_ctor_get(x_7, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_7); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___closed__1() { _start: { -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Name_hash___override(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_1, x_4, x_2); -return x_5; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___closed__2() { _start: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_5); -return x_6; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); +return x_1; } -else -{ -lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Name_hash___override(x_9); -x_12 = lean_uint64_to_usize(x_11); -x_13 = 1; -x_14 = lean_usize_sub(x_1, x_13); -x_15 = 5; -x_16 = lean_usize_mul(x_15, x_14); -x_17 = lean_usize_shift_right(x_12, x_16); -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_5, x_18); -lean_dec(x_5); -x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_6, x_17, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_19; -x_6 = x_20; -goto _start; } +static lean_object* _init_l_Lean_Meta_Grind_preprocessPattern___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_preprocessPattern___lambda__1), 6, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_Meta_Grind_preprocessPattern___closed__1; +x_12 = l_Lean_Meta_Grind_preprocessPattern___closed__2; lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_1, 1); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 0); -lean_dec(x_11); -x_12 = lean_array_push(x_5, x_3); -x_13 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_12); -return x_1; -} -else +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_13 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_9, x_11, x_12, x_3, x_4, x_5, x_6, x_10); +if (lean_obj_tag(x_13) == 0) { lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_1); -x_14 = lean_array_push(x_5, x_3); -x_15 = lean_array_push(x_6, x_4); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_preprocessPattern___closed__3; +if (x_2 == 0) +{ +lean_object* x_17; +x_17 = lean_apply_6(x_16, x_14, x_3, x_4, x_5, x_6, x_15); +return x_17; } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_name_eq(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_18; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_18 = lean_grind_normalize(x_14, x_3, x_4, x_5, x_6, x_15); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_19; lean_object* x_20; -lean_dec(x_6); -lean_dec(x_5); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_2, x_19); -lean_dec(x_2); -x_2 = x_20; -goto _start; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_apply_6(x_16, x_19, x_3, x_4, x_5, x_6, x_20); +return x_21; } else { uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -if (x_22 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 1); -lean_dec(x_23); -x_24 = lean_ctor_get(x_1, 0); -lean_dec(x_24); -x_25 = lean_array_fset(x_5, x_2, x_3); -x_26 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -lean_ctor_set(x_1, 1, x_26); -lean_ctor_set(x_1, 0, x_25); -return x_1; +return x_18; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +else +{ +uint8_t x_26; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) +{ +return x_13; } else { lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_1); -x_27 = lean_array_fset(x_5, x_2, x_3); -x_28 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); +x_27 = lean_ctor_get(x_13, 0); +x_28 = lean_ctor_get(x_13, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_13); x_29 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); @@ -1524,866 +1645,792 @@ return x_29; } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessPattern___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l_Lean_Meta_Grind_preprocessPattern(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1() { +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedOrigin() { _start: { lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); +x_1 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1() { _start: { -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = 1; -x_9 = 5; -x_10 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_14 = lean_nat_dec_lt(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.decl", 27, 27); return x_1; } -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_name_eq(x_4, x_19); -if (x_21 == 0) +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2() { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_free_object(x_15); -x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_array_fset(x_17, x_12, x_23); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_24); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3() { +_start: { -lean_object* x_25; -lean_dec(x_20); -lean_dec(x_19); -lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_4); -x_25 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_25); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4() { +_start: { -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_name_eq(x_4, x_26); -if (x_28 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5() { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_array_fset(x_17, x_12, x_30); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6() { +_start: { -lean_object* x_32; lean_object* x_33; -lean_dec(x_27); -lean_dec(x_26); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_4); -lean_ctor_set(x_32, 1, x_5); -x_33 = lean_array_fset(x_17, x_12, x_32); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_33); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.fvar", 27, 27); return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} } -case 1: +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8() { +_start: { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9() { +_start: { -lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -x_35 = lean_ctor_get(x_15, 0); -x_36 = lean_usize_shift_right(x_2, x_9); -x_37 = lean_usize_add(x_3, x_8); -x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_35, x_36, x_37, x_4, x_5); -lean_ctor_set(x_15, 0, x_38); -x_39 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.stx", 26, 26); return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10() { +_start: { -lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_15, 0); -lean_inc(x_40); -lean_dec(x_15); -x_41 = lean_usize_shift_right(x_2, x_9); -x_42 = lean_usize_add(x_3, x_8); -x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_40, x_41, x_42, x_4, x_5); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_array_fset(x_17, x_12, x_44); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -default: +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11() { +_start: { -lean_object* x_46; lean_object* x_47; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_4); -lean_ctor_set(x_46, 1, x_5); -x_47 = lean_array_fset(x_17, x_12, x_46); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_47); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.Origin.other", 28, 28); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13() { +_start: { -lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_1, 0); -lean_inc(x_48); -lean_dec(x_1); -x_49 = 1; -x_50 = 5; -x_51 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; -x_52 = lean_usize_land(x_2, x_51); -x_53 = lean_usize_to_nat(x_52); -x_54 = lean_array_get_size(x_48); -x_55 = lean_nat_dec_lt(x_53, x_54); -lean_dec(x_54); -if (x_55 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14() { +_start: { -lean_object* x_56; -lean_dec(x_53); -lean_dec(x_5); -lean_dec(x_4); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_48); -return x_56; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614_(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_array_fget(x_48, x_53); -x_58 = lean_box(0); -x_59 = lean_array_fset(x_48, x_53, x_58); -switch (lean_obj_tag(x_57)) { +switch (lean_obj_tag(x_1)) { case 0: { -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_57, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_62 = x_57; -} else { - lean_dec_ref(x_57); - x_62 = lean_box(0); -} -x_63 = lean_name_eq(x_4, x_60); -if (x_63 == 0) +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_unsigned_to_nat(1024u); +x_5 = lean_nat_dec_le(x_4, x_2); +x_6 = l_Lean_Name_reprPrec(x_3, x_4); +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +if (x_5 == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_array_fset(x_59, x_53, x_65); -lean_dec(x_53); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_10 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = 0; +x_12 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); +x_13 = l_Repr_addAppParen(x_12, x_2); +return x_13; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_61); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_68 = lean_alloc_ctor(0, 2, 0); -} else { - x_68 = x_62; -} -lean_ctor_set(x_68, 0, x_4); -lean_ctor_set(x_68, 1, x_5); -x_69 = lean_array_fset(x_59, x_53, x_68); -lean_dec(x_53); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; +lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; +x_15 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_8); +x_16 = 0; +x_17 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_16); +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; } } case 1: { -lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_71 = lean_ctor_get(x_57, 0); -lean_inc(x_71); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - x_72 = x_57; -} else { - lean_dec_ref(x_57); - x_72 = lean_box(0); -} -x_73 = lean_usize_shift_right(x_2, x_50); -x_74 = lean_usize_add(x_3, x_49); -x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_71, x_73, x_74, x_4, x_5); -if (lean_is_scalar(x_72)) { - x_76 = lean_alloc_ctor(1, 1, 0); -} else { - x_76 = x_72; -} -lean_ctor_set(x_76, 0, x_75); -x_77 = lean_array_fset(x_59, x_53, x_76); -lean_dec(x_53); -x_78 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_78, 0, x_77); -return x_78; -} -default: +lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_19 = lean_ctor_get(x_1, 0); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_unsigned_to_nat(1024u); +x_21 = lean_nat_dec_le(x_20, x_2); +x_22 = l_Lean_Name_reprPrec(x_19, x_20); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8; +x_24 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +if (x_21 == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_4); -lean_ctor_set(x_79, 1, x_5); -x_80 = lean_array_fset(x_59, x_53, x_79); -lean_dec(x_53); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; -} -} -} -} +lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_25 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_26 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = 0; +x_28 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); +x_29 = l_Repr_addAppParen(x_28, x_2); +return x_29; } else { -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) -{ -lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; -x_83 = lean_unsigned_to_nat(0u); -x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(x_1, x_83, x_4, x_5); -x_85 = 7; -x_86 = lean_usize_dec_le(x_85, x_3); -if (x_86 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; +x_31 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_24); +x_32 = 0; +x_33 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); +x_34 = l_Repr_addAppParen(x_33, x_2); +return x_34; +} +} +case 2: { -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_dec_lt(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) +uint8_t x_35; +x_35 = !lean_is_exclusive(x_1); +if (x_35 == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_84, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 1); -lean_inc(x_91); -lean_dec(x_84); -x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_3, x_90, x_91, lean_box(0), x_83, x_92); -lean_dec(x_91); -lean_dec(x_90); -return x_93; -} -else +lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_36 = lean_ctor_get(x_1, 0); +x_37 = lean_ctor_get(x_1, 1); +x_38 = lean_unsigned_to_nat(1024u); +x_39 = lean_nat_dec_le(x_38, x_2); +x_40 = l_Lean_Name_reprPrec(x_36, x_38); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11; +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_40); +lean_ctor_set(x_1, 0, x_41); +x_42 = lean_box(1); +x_43 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_43, 0, x_1); +lean_ctor_set(x_43, 1, x_42); +x_44 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_37, x_38); +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +if (x_39 == 0) { -return x_84; -} +lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; +x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_47 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +x_48 = 0; +x_49 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_48); +x_50 = l_Repr_addAppParen(x_49, x_2); +return x_50; } else { -return x_84; +lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_51 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; +x_52 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_45); +x_53 = 0; +x_54 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_53); +x_55 = l_Repr_addAppParen(x_54, x_2); +return x_55; } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; -x_94 = lean_ctor_get(x_1, 0); -x_95 = lean_ctor_get(x_1, 1); -lean_inc(x_95); -lean_inc(x_94); +lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_56 = lean_ctor_get(x_1, 0); +x_57 = lean_ctor_get(x_1, 1); +lean_inc(x_57); +lean_inc(x_56); lean_dec(x_1); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -x_97 = lean_unsigned_to_nat(0u); -x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(x_96, x_97, x_4, x_5); -x_99 = 7; -x_100 = lean_usize_dec_le(x_99, x_3); -if (x_100 == 0) -{ -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); -x_102 = lean_unsigned_to_nat(4u); -x_103 = lean_nat_dec_lt(x_101, x_102); -lean_dec(x_101); -if (x_103 == 0) +x_58 = lean_unsigned_to_nat(1024u); +x_59 = lean_nat_dec_le(x_58, x_2); +x_60 = l_Lean_Name_reprPrec(x_56, x_58); +x_61 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11; +x_62 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_60); +x_63 = lean_box(1); +x_64 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +x_65 = l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2170_(x_57, x_58); +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +if (x_59 == 0) { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_98, 1); -lean_inc(x_105); -lean_dec(x_98); -x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; +lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; +x_67 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_68 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_66); +x_69 = 0; +x_70 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set_uint8(x_70, sizeof(void*)*1, x_69); +x_71 = l_Repr_addAppParen(x_70, x_2); +return x_71; } else { -return x_98; +lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; +x_72 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; +x_73 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_66); +x_74 = 0; +x_75 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set_uint8(x_75, sizeof(void*)*1, x_74); +x_76 = l_Repr_addAppParen(x_75, x_2); +return x_76; } } -else +} +default: { -return x_98; +lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_77 = lean_ctor_get(x_1, 0); +lean_inc(x_77); +lean_dec(x_1); +x_78 = lean_unsigned_to_nat(1024u); +x_79 = lean_nat_dec_le(x_78, x_2); +x_80 = l_Lean_Name_reprPrec(x_77, x_78); +x_81 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14; +x_82 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_80); +if (x_79 == 0) +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; +x_83 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4; +x_84 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_82); +x_85 = 0; +x_86 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set_uint8(x_86, sizeof(void*)*1, x_85); +x_87 = l_Repr_addAppParen(x_86, x_2); +return x_87; +} +else +{ +lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; +x_88 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5; +x_89 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_82); +x_90 = 0; +x_91 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set_uint8(x_91, sizeof(void*)*1, x_90); +x_92 = l_Repr_addAppParen(x_91, x_2); +return x_92; } } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Name_hash___override(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_1, x_5, x_6, x_2, x_3); -return x_7; +lean_object* x_3; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614_(x_1, x_2); +lean_dec(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.EMatchTheorem", 36, 36); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_instReprOrigin() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.EMatchTheorems.insert", 37, 37); +x_1 = l_Lean_Meta_Grind_instReprOrigin___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3() { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811_(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} +switch (lean_obj_tag(x_1)) { +case 0: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_name_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; } -static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; -x_2 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2; -x_3 = lean_unsigned_to_nat(64u); -x_4 = lean_unsigned_to_nat(6u); -x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +uint8_t x_6; +lean_dec(x_2); +lean_dec(x_1); +x_6 = 0; return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object* x_1, lean_object* x_2) { -_start: +case 1: { -lean_object* x_3; -x_3 = lean_ctor_get(x_2, 4); -lean_inc(x_3); -if (lean_obj_tag(x_3) == 0) +if (lean_obj_tag(x_2) == 1) { -lean_object* x_4; lean_object* x_5; -lean_dec(x_2); +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); lean_dec(x_1); -x_4 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; -x_5 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_4); -return x_5; +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_name_eq(x_7, x_8); +lean_dec(x_8); +lean_dec(x_7); +return x_9; } else { -lean_object* x_6; -x_6 = lean_ctor_get(x_3, 0); -lean_inc(x_6); -if (lean_obj_tag(x_6) == 2) -{ -uint8_t x_7; -x_7 = !lean_is_exclusive(x_2); -if (x_7 == 0) +uint8_t x_10; +lean_dec(x_2); +lean_dec(x_1); +x_10 = 0; +return x_10; +} +} +case 2: { -lean_object* x_8; uint8_t x_9; -x_8 = lean_ctor_get(x_2, 4); -lean_dec(x_8); -x_9 = !lean_is_exclusive(x_3); -if (x_9 == 0) +if (lean_obj_tag(x_2) == 2) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_3, 1); -x_11 = lean_ctor_get(x_3, 0); -lean_dec(x_11); -x_12 = lean_ctor_get(x_6, 0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); lean_inc(x_12); -lean_dec(x_6); -lean_ctor_set(x_2, 4, x_10); -lean_inc(x_1); -x_13 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_12); -if (lean_obj_tag(x_13) == 0) +lean_dec(x_1); +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +lean_dec(x_2); +x_15 = lean_name_eq(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +if (x_15 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_box(0); -lean_ctor_set(x_3, 1, x_14); -lean_ctor_set(x_3, 0, x_2); -x_15 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_12, x_3); -return x_15; +uint8_t x_16; +lean_dec(x_14); +lean_dec(x_12); +x_16 = 0; +return x_16; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 0); -lean_inc(x_16); -lean_dec(x_13); -lean_ctor_set(x_3, 1, x_16); -lean_ctor_set(x_3, 0, x_2); -x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_12, x_3); +uint8_t x_17; +x_17 = l_Lean_Syntax_structEq(x_12, x_14); return x_17; } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_3, 1); -lean_inc(x_18); -lean_dec(x_3); -x_19 = lean_ctor_get(x_6, 0); -lean_inc(x_19); -lean_dec(x_6); -lean_ctor_set(x_2, 4, x_18); -lean_inc(x_1); -x_20 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_19); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_2); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_19, x_22); -return x_23; +uint8_t x_18; +lean_dec(x_2); +lean_dec(x_1); +x_18 = 0; +return x_18; } -else +} +default: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 0); -lean_inc(x_24); +if (lean_obj_tag(x_2) == 3) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_1, 0); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_ctor_get(x_2, 0); +lean_inc(x_20); +lean_dec(x_2); +x_21 = lean_name_eq(x_19, x_20); lean_dec(x_20); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_2); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_19, x_25); -return x_26; -} -} +lean_dec(x_19); +return x_21; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_27 = lean_ctor_get(x_2, 0); -x_28 = lean_ctor_get(x_2, 1); -x_29 = lean_ctor_get(x_2, 2); -x_30 = lean_ctor_get(x_2, 3); -x_31 = lean_ctor_get(x_2, 5); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); +uint8_t x_22; lean_dec(x_2); -x_32 = lean_ctor_get(x_3, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_3)) { - lean_ctor_release(x_3, 0); - lean_ctor_release(x_3, 1); - x_33 = x_3; -} else { - lean_dec_ref(x_3); - x_33 = lean_box(0); +lean_dec(x_1); +x_22 = 0; +return x_22; } -x_34 = lean_ctor_get(x_6, 0); -lean_inc(x_34); -lean_dec(x_6); -x_35 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_35, 0, x_27); -lean_ctor_set(x_35, 1, x_28); -lean_ctor_set(x_35, 2, x_29); -lean_ctor_set(x_35, 3, x_30); -lean_ctor_set(x_35, 4, x_32); -lean_ctor_set(x_35, 5, x_31); -lean_inc(x_1); -x_36 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_34); -if (lean_obj_tag(x_36) == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_box(0); -if (lean_is_scalar(x_33)) { - x_38 = lean_alloc_ctor(1, 2, 0); -} else { - x_38 = x_33; } -lean_ctor_set(x_38, 0, x_35); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_34, x_38); -return x_39; } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811____boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_36, 0); -lean_inc(x_40); -lean_dec(x_36); -if (lean_is_scalar(x_33)) { - x_41 = lean_alloc_ctor(1, 2, 0); -} else { - x_41 = x_33; +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811_(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; } -lean_ctor_set(x_41, 0, x_35); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_34, x_41); -return x_42; } +static lean_object* _init_l_Lean_Meta_Grind_instBEqOrigin___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_811____boxed), 2, 0); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind_instBEqOrigin() { +_start: { -lean_object* x_43; lean_object* x_44; -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_43 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; -x_44 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_43); -return x_44; +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instBEqOrigin___closed__1; +return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_key___boxed(lean_object* x_1) { _start: { -lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_2); +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Origin_key(x_1); lean_dec(x_1); -return x_6; +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_1, x_4, x_3); +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); lean_dec(x_3); -return x_5; +x_5 = l_Lean_MessageData_ofConst(x_2); +x_6 = lean_apply_2(x_4, lean_box(0), x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_1, x_2); -lean_dec(x_2); -return x_3; +switch (lean_obj_tag(x_4)) { +case 0: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_inc(x_1); +x_7 = l_Lean_mkConstWithLevelParams___rarg(x_1, x_2, x_3, x_5); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg___lambda__1), 2, 1); +lean_closure_set(x_8, 0, x_1); +x_9 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_8); +return x_9; } +case 1: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_3); +lean_dec(x_2); +x_10 = lean_ctor_get(x_4, 0); +lean_inc(x_10); +lean_dec(x_4); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_dec(x_1); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_Expr_fvar___override(x_10); +x_14 = l_Lean_MessageData_ofExpr(x_13); +x_15 = lean_apply_2(x_12, lean_box(0), x_14); +return x_15; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +case 2: { -size_t x_7; lean_object* x_8; -x_7 = lean_unbox_usize(x_1); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_ctor_get(x_4, 1); +lean_inc(x_16); +lean_dec(x_4); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_7, x_2, x_3, x_4, x_5, x_6); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_MessageData_ofSyntax(x_16); +x_20 = lean_apply_2(x_18, lean_box(0), x_19); +return x_20; +} +default: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_3); lean_dec(x_2); -return x_8; +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +lean_dec(x_4); +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +lean_dec(x_1); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_MessageData_ofName(x_21); +x_25 = lean_apply_2(x_23, lean_box(0), x_24); +return x_25; +} } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp(lean_object* x_1) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_1, x_6, x_7, x_4, x_5); -return x_8; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_Origin_pp___rarg), 4, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqOrigin__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -x_8 = l_Lean_Expr_isConst(x_7); -if (x_8 == 0) -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -lean_dec(x_1); -x_10 = l_Array_isEmpty___rarg(x_9); -if (x_10 == 0) -{ -size_t x_11; size_t x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_array_size(x_9); -x_12 = 0; -lean_inc(x_9); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_11, x_12, x_9, x_2, x_3, x_4, x_5, x_6); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_13, 0); -x_16 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_15); -lean_dec(x_7); -lean_ctor_set(x_13, 0, x_16); -return x_13; +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_Meta_Grind_Origin_key(x_1); +x_4 = l_Lean_Meta_Grind_Origin_key(x_2); +x_5 = lean_name_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqOrigin__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_13, 0); -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_13); -x_19 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_17); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_instBEqOrigin__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } } -else +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableOrigin(lean_object* x_1) { +_start: { -lean_object* x_21; -lean_dec(x_9); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_7); -lean_ctor_set(x_21, 1, x_6); -return x_21; +lean_object* x_2; uint64_t x_3; +x_2 = l_Lean_Meta_Grind_Origin_key(x_1); +x_3 = l_Lean_Name_hash___override(x_2); +lean_dec(x_2); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableOrigin___boxed(lean_object* x_1) { +_start: { -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_1, 0); -lean_inc(x_22); +uint64_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_instHashableOrigin(x_1); lean_dec(x_1); -x_23 = l_Array_isEmpty___rarg(x_22); -if (x_23 == 0) -{ -size_t x_24; size_t x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_array_size(x_22); -x_25 = 0; -lean_inc(x_22); -x_26 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_24, x_25, x_22, x_2, x_3, x_4, x_5, x_6); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_26, 0); -x_29 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_28); -lean_dec(x_7); -lean_ctor_set(x_26, 0, x_29); -return x_26; +x_3 = lean_box_uint64(x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1() { +_start: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_26, 0); -x_31 = lean_ctor_get(x_26, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_26); -x_32 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_30); -lean_dec(x_7); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -return x_33; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -else +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2() { +_start: { -lean_object* x_34; lean_object* x_35; -lean_dec(x_22); -x_34 = l_Lean_Expr_constName_x21(x_7); -lean_inc(x_34); -x_35 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_34, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_35) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Expr_bvar___override(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3() { +_start: { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_37 = lean_ctor_get(x_35, 0); -x_38 = lean_ctor_get(x_35, 1); -x_39 = l_Lean_ConstantInfo_levelParams(x_37); -lean_dec(x_37); -x_40 = l_List_isEmpty___rarg(x_39); -lean_dec(x_39); -if (x_40 == 0) -{ -lean_object* x_41; -lean_free_object(x_35); -lean_dec(x_7); -x_41 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_38); -return x_41; -} -else -{ -lean_dec(x_34); -lean_ctor_set(x_35, 0, x_7); -return x_35; -} -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_42 = lean_ctor_get(x_35, 0); -x_43 = lean_ctor_get(x_35, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_35); -x_44 = l_Lean_ConstantInfo_levelParams(x_42); -lean_dec(x_42); -x_45 = l_List_isEmpty___rarg(x_44); -lean_dec(x_44); -if (x_45 == 0) -{ -lean_object* x_46; -lean_dec(x_7); -x_46 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_43); -return x_46; -} -else -{ -lean_object* x_47; -lean_dec(x_34); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_7); -lean_ctor_set(x_47, 1, x_43); -return x_47; -} -} -} -else -{ -uint8_t x_48; -lean_dec(x_34); -lean_dec(x_7); -x_48 = !lean_is_exclusive(x_35); -if (x_48 == 0) -{ -return x_35; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_35, 0); -x_50 = lean_ctor_get(x_35, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_35); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; -} -} -} -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1; +x_3 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l_Lean_Meta_Grind_instInhabitedOrigin___closed__1; +x_6 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set(x_6, 3, x_1); +lean_ctor_set(x_6, 4, x_1); +lean_ctor_set(x_6, 5, x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem() { _start: { -lean_object* x_7; -x_7 = l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_7; +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3; +return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1() { _start: { lean_object* x_1; @@ -2391,15772 +2438,27350 @@ x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1() { _start: { lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorems() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private", 8, 8); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2() { +static lean_object* _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); +x_1 = l_Id_instMonad; +x_2 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3() { +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_dec(x_5); +return x_6; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_Origin_key(x_9); +x_12 = l_Lean_Name_hash___override(x_11); +lean_dec(x_11); +x_13 = lean_uint64_to_usize(x_12); +x_14 = 1; +x_15 = lean_usize_sub(x_1, x_14); +x_16 = 5; +x_17 = lean_usize_mul(x_16, x_15); +x_18 = lean_usize_shift_right(x_13, x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_5, x_19); +lean_dec(x_5); +x_21 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_6, x_18, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_20; +x_6 = x_21; +goto _start; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5() { +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Meta", 4, 4); +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); return x_1; } -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Tactic", 6, 6); -return x_1; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_Origin_key(x_3); +x_19 = l_Lean_Meta_Grind_Origin_key(x_17); +lean_dec(x_17); +x_20 = lean_name_eq(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_6); +lean_dec(x_5); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_2, x_21); +lean_dec(x_2); +x_2 = x_22; +goto _start; } +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_1); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_1, 1); +lean_dec(x_25); +x_26 = lean_ctor_get(x_1, 0); +lean_dec(x_26); +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_28); +lean_ctor_set(x_1, 0, x_27); +return x_1; } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_1); +x_29 = lean_array_fset(x_5, x_2, x_3); +x_30 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9() { +} +static size_t _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = 5; +x_3 = lean_usize_shift_left(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10() { +static size_t _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1; +x_3 = lean_usize_sub(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11() { +static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("EMatchTheorem", 13, 13); +x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13() { -_start: +if (lean_obj_tag(x_1) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Name_num___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14() { -_start: +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15() { -_start: +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17() { -_start: +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ematchTheoremsExt", 17, 17); +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = l_Lean_Meta_Grind_Origin_key(x_4); +x_22 = l_Lean_Meta_Grind_Origin_key(x_19); +x_23 = lean_name_eq(x_21, x_22); +lean_dec(x_22); +lean_dec(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_free_object(x_15); +x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_array_fset(x_17, x_12, x_25); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_26); return x_1; } -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_27; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_27 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_27); +return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_EMatchTheorems_insert), 2, 0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_28 = lean_ctor_get(x_15, 0); +x_29 = lean_ctor_get(x_15, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_15); +x_30 = l_Lean_Meta_Grind_Origin_key(x_4); +x_31 = l_Lean_Meta_Grind_Origin_key(x_28); +x_32 = lean_name_eq(x_30, x_31); +lean_dec(x_31); +lean_dec(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_28, x_29, x_4, x_5); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +x_35 = lean_array_fset(x_17, x_12, x_34); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_35); return x_1; } -} -static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +lean_object* x_36; lean_object* x_37; +lean_dec(x_29); +lean_dec(x_28); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_4); +lean_ctor_set(x_36, 1, x_5); +x_37 = lean_array_fset(x_17, x_12, x_36); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_37); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18; -x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19; -x_4 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1; -x_5 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20; -x_6 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_6, 0, x_2); -lean_ctor_set(x_6, 1, x_3); -lean_ctor_set(x_6, 2, x_4); -lean_ctor_set(x_6, 3, x_5); -x_7 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_6, x_1); -return x_7; -} } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1() { -_start: +case 1: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Eq", 2, 2); +uint8_t x_38; +x_38 = !lean_is_exclusive(x_15); +if (x_38 == 0) +{ +lean_object* x_39; size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; +x_39 = lean_ctor_get(x_15, 0); +x_40 = lean_usize_shift_right(x_2, x_9); +x_41 = lean_usize_add(x_3, x_8); +x_42 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_39, x_40, x_41, x_4, x_5); +lean_ctor_set(x_15, 0, x_42); +x_43 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_43); return x_1; } -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_44; size_t x_45; size_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_ctor_get(x_15, 0); +lean_inc(x_44); +lean_dec(x_15); +x_45 = lean_usize_shift_right(x_2, x_9); +x_46 = lean_usize_add(x_3, x_8); +x_47 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_44, x_45, x_46, x_4, x_5); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = lean_array_fset(x_17, x_12, x_48); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_49); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3() { -_start: +default: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("HEq", 3, 3); +lean_object* x_50; lean_object* x_51; +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_4); +lean_ctor_set(x_50, 1, x_5); +x_51 = lean_array_fset(x_17, x_12, x_50); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_51); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Iff", 3, 3); -return x_1; -} +lean_object* x_52; size_t x_53; size_t x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_52 = lean_ctor_get(x_1, 0); +lean_inc(x_52); +lean_dec(x_1); +x_53 = 1; +x_54 = 5; +x_55 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_56 = lean_usize_land(x_2, x_55); +x_57 = lean_usize_to_nat(x_56); +x_58 = lean_array_get_size(x_52); +x_59 = lean_nat_dec_lt(x_57, x_58); +lean_dec(x_58); +if (x_59 == 0) +{ +lean_object* x_60; +lean_dec(x_57); +lean_dec(x_5); +lean_dec(x_4); +x_60 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_60, 0, x_52); +return x_60; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_array_fget(x_52, x_57); +x_62 = lean_box(0); +x_63 = lean_array_fset(x_52, x_57, x_62); +switch (lean_obj_tag(x_61)) { +case 0: +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_64 = lean_ctor_get(x_61, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_61, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_66 = x_61; +} else { + lean_dec_ref(x_61); + x_66 = lean_box(0); } +x_67 = l_Lean_Meta_Grind_Origin_key(x_4); +x_68 = l_Lean_Meta_Grind_Origin_key(x_64); +x_69 = lean_name_eq(x_67, x_68); +lean_dec(x_68); +lean_dec(x_67); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_dec(x_66); +x_70 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_64, x_65, x_4, x_5); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_72 = lean_array_fset(x_63, x_57, x_71); +lean_dec(x_57); +x_73 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_73, 0, x_72); +return x_73; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("And", 3, 3); -return x_1; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_65); +lean_dec(x_64); +if (lean_is_scalar(x_66)) { + x_74 = lean_alloc_ctor(0, 2, 0); +} else { + x_74 = x_66; +} +lean_ctor_set(x_74, 0, x_4); +lean_ctor_set(x_74, 1, x_5); +x_75 = lean_array_fset(x_63, x_57, x_74); +lean_dec(x_57); +x_76 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_76, 0, x_75); +return x_76; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8() { -_start: +case 1: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_77; lean_object* x_78; size_t x_79; size_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_77 = lean_ctor_get(x_61, 0); +lean_inc(x_77); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + x_78 = x_61; +} else { + lean_dec_ref(x_61); + x_78 = lean_box(0); } +x_79 = lean_usize_shift_right(x_2, x_54); +x_80 = lean_usize_add(x_3, x_53); +x_81 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_77, x_79, x_80, x_4, x_5); +if (lean_is_scalar(x_78)) { + x_82 = lean_alloc_ctor(1, 1, 0); +} else { + x_82 = x_78; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9() { -_start: +lean_ctor_set(x_82, 0, x_81); +x_83 = lean_array_fset(x_63, x_57, x_82); +lean_dec(x_57); +x_84 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_84, 0, x_83); +return x_84; +} +default: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Or", 2, 2); -return x_1; +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_4); +lean_ctor_set(x_85, 1, x_5); +x_86 = lean_array_fset(x_63, x_57, x_85); +lean_dec(x_57); +x_87 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_87, 0, x_86); +return x_87; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Not", 3, 3); -return x_1; +uint8_t x_88; +x_88 = !lean_is_exclusive(x_1); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; size_t x_91; uint8_t x_92; +x_89 = lean_unsigned_to_nat(0u); +x_90 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_1, x_89, x_4, x_5); +x_91 = 7; +x_92 = lean_usize_dec_le(x_91, x_3); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_93 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_90); +x_94 = lean_unsigned_to_nat(4u); +x_95 = lean_nat_dec_lt(x_93, x_94); +lean_dec(x_93); +if (x_95 == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_90, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_90, 1); +lean_inc(x_97); +lean_dec(x_90); +x_98 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3; +x_99 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_3, x_96, x_97, lean_box(0), x_89, x_98); +lean_dec(x_97); +lean_dec(x_96); +return x_99; } +else +{ +return x_90; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +return x_90; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14() { -_start: +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; size_t x_105; uint8_t x_106; +x_100 = lean_ctor_get(x_1, 0); +x_101 = lean_ctor_get(x_1, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_1); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +x_103 = lean_unsigned_to_nat(0u); +x_104 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__5(x_102, x_103, x_4, x_5); +x_105 = 7; +x_106 = lean_usize_dec_le(x_105, x_3); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_107 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_104); +x_108 = lean_unsigned_to_nat(4u); +x_109 = lean_nat_dec_lt(x_107, x_108); +lean_dec(x_107); +if (x_109 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_110 = lean_ctor_get(x_104, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_104, 1); +lean_inc(x_111); +lean_dec(x_104); +x_112 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3; +x_113 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_3, x_110, x_111, lean_box(0), x_103, x_112); +lean_dec(x_111); +lean_dec(x_110); +return x_113; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +return x_104; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} +return x_104; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; -x_2 = lean_array_mk(x_1); -return x_2; +lean_object* x_4; uint64_t x_5; size_t x_6; size_t x_7; lean_object* x_8; +x_4 = l_Lean_Meta_Grind_Origin_key(x_2); +x_5 = l_Lean_Name_hash___override(x_4); +lean_dec(x_4); +x_6 = lean_uint64_to_usize(x_5); +x_7 = 1; +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_1, x_6, x_7, x_2, x_3); +return x_8; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames() { +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; -return x_1; +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_3); +x_6 = lean_box(0); +return x_6; } +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = lean_array_fget(x_1, x_3); +x_8 = l_Lean_Meta_Grind_Origin_key(x_7); +lean_dec(x_7); +x_9 = l_Lean_Meta_Grind_Origin_key(x_2); +x_10 = lean_name_eq(x_8, x_9); +lean_dec(x_9); +lean_dec(x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_3, x_11); +lean_dec(x_3); +x_3 = x_12; +goto _start; } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(lean_object* x_1) { -_start: +else { -lean_object* x_2; uint8_t x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; -x_3 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_2, x_1); -return x_3; +lean_object* x_14; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_3); +return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("[grind_dontcare]", 16, 16); +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_Origin_key(x_3); +x_13 = l_Lean_Meta_Grind_Origin_key(x_11); +lean_dec(x_11); +x_14 = lean_name_eq(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_dec(x_8); +lean_dec(x_4); return x_1; } -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} +uint8_t x_15; +x_15 = !lean_is_exclusive(x_1); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_1, 0); +lean_dec(x_16); +x_17 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_17); +return x_1; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; +lean_object* x_18; lean_object* x_19; +lean_dec(x_1); +x_18 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +x_19 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_19, 0, x_18); +return x_19; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare() { -_start: +} +case 1: { -lean_object* x_1; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3; +uint8_t x_20; +x_20 = !lean_is_exclusive(x_1); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_1, 0); +lean_dec(x_21); +x_22 = !lean_is_exclusive(x_10); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; size_t x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_10, 0); +x_24 = lean_array_set(x_4, x_8, x_9); +x_25 = lean_usize_shift_right(x_2, x_5); +x_26 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_23, x_25, x_3); +lean_inc(x_26); +x_27 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_26); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; +lean_ctor_set(x_10, 0, x_26); +x_28 = lean_array_set(x_24, x_8, x_10); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_28); return x_1; } +else +{ +lean_object* x_29; uint8_t x_30; +lean_dec(x_26); +lean_free_object(x_10); +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; +x_31 = lean_array_set(x_24, x_8, x_29); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_31); +return x_1; } -static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_29, 0); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_29); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_array_set(x_24, x_8, x_34); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_35); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__2() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ground_pat", 10, 10); +lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; +x_36 = lean_ctor_get(x_10, 0); +lean_inc(x_36); +lean_dec(x_10); +x_37 = lean_array_set(x_4, x_8, x_9); +x_38 = lean_usize_shift_right(x_2, x_5); +x_39 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_36, x_38, x_3); +lean_inc(x_39); +x_40 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_39); +x_42 = lean_array_set(x_37, x_8, x_41); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_42); return x_1; } -} -static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; -x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_39); +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +lean_dec(x_40); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_43)) { + lean_ctor_release(x_43, 0); + lean_ctor_release(x_43, 1); + x_46 = x_43; +} else { + lean_dec_ref(x_43); + x_46 = lean_box(0); +} +if (lean_is_scalar(x_46)) { + x_47 = lean_alloc_ctor(0, 2, 0); +} else { + x_47 = x_46; } +lean_ctor_set(x_47, 0, x_44); +lean_ctor_set(x_47, 1, x_45); +x_48 = lean_array_set(x_37, x_8, x_47); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_48); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; -x_3 = l_Lean_mkAnnotation(x_2, x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; -x_3 = l_Lean_annotation_x3f(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; size_t x_52; lean_object* x_53; lean_object* x_54; lean_dec(x_1); -return x_2; -} +x_49 = lean_ctor_get(x_10, 0); +lean_inc(x_49); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_50 = x_10; +} else { + lean_dec_ref(x_10); + x_50 = lean_box(0); } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); -if (lean_obj_tag(x_2) == 0) +x_51 = lean_array_set(x_4, x_8, x_9); +x_52 = lean_usize_shift_right(x_2, x_5); +x_53 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_49, x_52, x_3); +lean_inc(x_53); +x_54 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_53); +if (lean_obj_tag(x_54) == 0) { -uint8_t x_3; -x_3 = 0; -return x_3; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +if (lean_is_scalar(x_50)) { + x_55 = lean_alloc_ctor(1, 1, 0); +} else { + x_55 = x_50; +} +lean_ctor_set(x_55, 0, x_53); +x_56 = lean_array_set(x_51, x_8, x_55); +lean_dec(x_8); +x_57 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_57, 0, x_56); +return x_57; } else { -uint8_t x_4; -lean_dec(x_2); -x_4 = 1; -return x_4; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_53); +lean_dec(x_50); +x_58 = lean_ctor_get(x_54, 0); +lean_inc(x_58); +lean_dec(x_54); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_61 = x_58; +} else { + lean_dec_ref(x_58); + x_61 = lean_box(0); } +if (lean_is_scalar(x_61)) { + x_62 = lean_alloc_ctor(0, 2, 0); +} else { + x_62 = x_61; } +lean_ctor_set(x_62, 0, x_59); +lean_ctor_set(x_62, 1, x_60); +x_63 = lean_array_set(x_51, x_8, x_62); +lean_dec(x_8); +x_64 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_64, 0, x_63); +return x_64; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; } } -LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isPatternDontCare(lean_object* x_1) { -_start: +default: { -lean_object* x_2; uint8_t x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_3 = lean_expr_eqv(x_1, x_2); -return x_3; -} +lean_dec(x_8); +lean_dec(x_4); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isPatternDontCare___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_isPatternDontCare(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = l_Lean_Expr_isBVar(x_1); -if (x_2 == 0) +else { -lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_4 = lean_expr_eqv(x_1, x_3); -if (x_4 == 0) +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_65 = lean_ctor_get(x_1, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_1, 1); +lean_inc(x_66); +x_67 = lean_unsigned_to_nat(0u); +x_68 = l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(x_65, x_3, x_67); +if (lean_obj_tag(x_68) == 0) { -uint8_t x_5; -x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); -return x_5; +lean_dec(x_66); +lean_dec(x_65); +return x_1; } else { -uint8_t x_6; -x_6 = 1; -return x_6; -} +uint8_t x_69; +x_69 = !lean_is_exclusive(x_1); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_70 = lean_ctor_get(x_1, 1); +lean_dec(x_70); +x_71 = lean_ctor_get(x_1, 0); +lean_dec(x_71); +x_72 = lean_ctor_get(x_68, 0); +lean_inc(x_72); +lean_dec(x_68); +lean_inc(x_72); +x_73 = l_Array_eraseIdx___rarg(x_65, x_72, lean_box(0)); +x_74 = l_Array_eraseIdx___rarg(x_66, x_72, lean_box(0)); +lean_ctor_set(x_1, 1, x_74); +lean_ctor_set(x_1, 0, x_73); +return x_1; } else { -uint8_t x_7; -x_7 = 1; -return x_7; +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_1); +x_75 = lean_ctor_get(x_68, 0); +lean_inc(x_75); +lean_dec(x_68); +lean_inc(x_75); +x_76 = l_Array_eraseIdx___rarg(x_65, x_75, lean_box(0)); +x_77 = l_Array_eraseIdx___rarg(x_66, x_75, lean_box(0)); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked(" ", 1, 1); -return x_1; +lean_object* x_3; uint64_t x_4; size_t x_5; lean_object* x_6; +x_3 = l_Lean_Meta_Grind_Origin_key(x_2); +x_4 = l_Lean_Name_hash___override(x_3); +lean_dec(x_3); +x_5 = lean_uint64_to_usize(x_4); +x_6 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_1, x_5, x_2); +return x_6; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3() { -_start: +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; -} +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; -x_5 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_4); -x_6 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_1); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1() { -_start: +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_name_eq(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed), 3, 0); -return x_1; -} +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("(", 1, 1); -return x_1; +lean_object* x_14; lean_object* x_15; +x_14 = lean_array_fget(x_2, x_4); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(")", 1, 1); -return x_1; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +if (lean_obj_tag(x_1) == 0) { -return x_6; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); +lean_dec(x_5); +switch (lean_obj_tag(x_11)) { +case 0: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_name_eq(x_3, x_12); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); +return x_1; +} +} +case 1: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; +goto _start; } -else +default: { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); +return x_19; +} } } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { +case 0: { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_name_eq(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +case 1: { -return x_6; +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; } -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +default: { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_35; +x_35 = lean_box(0); +return x_35; } -else -{ -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; } } -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; } } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10(x_1, x_4, x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) { +lean_dec(x_5); return x_6; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; -} -else -{ -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; +lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Name_hash___override(x_9); +x_12 = lean_uint64_to_usize(x_11); +x_13 = 1; +x_14 = lean_usize_sub(x_1, x_13); +x_15 = 5; +x_16 = lean_usize_mul(x_15, x_14); +x_17 = lean_usize_shift_right(x_12, x_16); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_5, x_18); +lean_dec(x_5); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(x_6, x_17, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_19; +x_6 = x_20; goto _start; } } -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) { -return x_6; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} } else { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_name_eq(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_6); +lean_dec(x_5); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_2); +x_2 = x_20; goto _start; } -} else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 1); lean_dec(x_23); -return x_24; +x_24 = lean_ctor_get(x_1, 0); +lean_dec(x_24); +x_25 = lean_array_fset(x_5, x_2, x_3); +x_26 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); +return x_1; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +if (lean_obj_tag(x_1) == 0) { -return x_6; +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_name_eq(x_4, x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_array_fset(x_17, x_12, x_23); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_24); +return x_1; } else { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +lean_object* x_25; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_25 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_25); +return x_1; } } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_name_eq(x_4, x_26); +if (x_28 == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = lean_array_fset(x_17, x_12, x_30); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_31); +return x_1; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} +lean_object* x_32; lean_object* x_33; +lean_dec(x_27); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_5); +x_33 = lean_array_fset(x_17, x_12, x_32); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_33); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; } -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +case 1: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_usize_shift_right(x_2, x_9); +x_37 = lean_usize_add(x_3, x_8); +x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(x_35, x_36, x_37, x_4, x_5); +lean_ctor_set(x_15, 0, x_38); +x_39 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); +return x_1; } else { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} +lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_15, 0); +lean_inc(x_40); +lean_dec(x_15); +x_41 = lean_usize_shift_right(x_2, x_9); +x_42 = lean_usize_add(x_3, x_8); +x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(x_40, x_41, x_42, x_4, x_5); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_array_fset(x_17, x_12, x_44); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_45); +return x_1; } -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; } -else +default: { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} +lean_object* x_46; lean_object* x_47; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_4); +lean_ctor_set(x_46, 1, x_5); +x_47 = lean_array_fset(x_17, x_12, x_46); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_47); +return x_1; } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: +else { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +lean_dec(x_1); +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_52 = lean_usize_land(x_2, x_51); +x_53 = lean_usize_to_nat(x_52); +x_54 = lean_array_get_size(x_48); +x_55 = lean_nat_dec_lt(x_53, x_54); +lean_dec(x_54); +if (x_55 == 0) { -return x_6; +lean_object* x_56; +lean_dec(x_53); +lean_dec(x_5); +lean_dec(x_4); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_array_fget(x_48, x_53); +x_58 = lean_box(0); +x_59 = lean_array_fset(x_48, x_53, x_58); +switch (lean_obj_tag(x_57)) { +case 0: { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_62 = x_57; +} else { + lean_dec_ref(x_57); + x_62 = lean_box(0); } -else +x_63 = lean_name_eq(x_4, x_60); +if (x_63 == 0) { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; -} +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_66 = lean_array_fset(x_59, x_53, x_65); +lean_dec(x_53); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_61); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_62; } -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +lean_ctor_set(x_68, 0, x_4); +lean_ctor_set(x_68, 1, x_5); +x_69 = lean_array_fset(x_59, x_53, x_68); +lean_dec(x_53); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; } } +case 1: +{ +lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_71 = lean_ctor_get(x_57, 0); +lean_inc(x_71); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_72 = x_57; +} else { + lean_dec_ref(x_57); + x_72 = lean_box(0); } +x_73 = lean_usize_shift_right(x_2, x_50); +x_74 = lean_usize_add(x_3, x_49); +x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(x_71, x_73, x_74, x_4, x_5); +if (lean_is_scalar(x_72)) { + x_76 = lean_alloc_ctor(1, 1, 0); +} else { + x_76 = x_72; } +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_array_fset(x_59, x_53, x_76); +lean_dec(x_53); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +default: { -return x_6; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_4); +lean_ctor_set(x_79, 1, x_5); +x_80 = lean_array_fset(x_59, x_53, x_79); +lean_dec(x_53); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} +} +} +} } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; +x_83 = lean_unsigned_to_nat(0u); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__15(x_1, x_83, x_4, x_5); +x_85 = 7; +x_86 = lean_usize_dec_le(x_85, x_3); +if (x_86 == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_dec_lt(x_87, x_88); +lean_dec(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_84, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 1); +lean_inc(x_91); +lean_dec(x_84); +x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } else { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +return x_84; } } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +return x_84; +} +} +else { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_1); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_unsigned_to_nat(0u); +x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__15(x_96, x_97, x_4, x_5); +x_99 = 7; +x_100 = lean_usize_dec_le(x_99, x_3); +if (x_100 == 0) +{ +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_dec_lt(x_101, x_102); +lean_dec(x_101); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +return x_98; +} +} +else +{ +return x_98; } } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Name_hash___override(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(x_1, x_5, x_6, x_2, x_3); +return x_7; } -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; } -else +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1() { +_start: { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.EMatchTheorem", 36, 36); +return x_1; } } -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2() { +_start: { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.EMatchTheorems.insert", 37, 37); +return x_1; } -else -{ -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; } +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2; +x_3 = lean_unsigned_to_nat(119u); +x_4 = lean_unsigned_to_nat(6u); +x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +lean_object* x_3; +x_3 = lean_ctor_get(x_2, 4); +lean_inc(x_3); +if (lean_obj_tag(x_3) == 0) { -return x_6; +lean_object* x_4; lean_object* x_5; +lean_dec(x_2); +lean_dec(x_1); +x_4 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; +x_5 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_4); +return x_5; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); -lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); -lean_dec(x_8); -if (x_11 == 0) +lean_object* x_6; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +if (lean_obj_tag(x_6) == 2) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +uint8_t x_7; +x_7 = !lean_is_exclusive(x_2); +if (x_7 == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_2, 5); +x_9 = lean_ctor_get(x_2, 4); +lean_dec(x_9); +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_3, 1); +x_12 = lean_ctor_get(x_3, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_6, 0); +lean_inc(x_13); +lean_dec(x_6); +lean_inc(x_8); +lean_ctor_set(x_2, 4, x_11); +x_14 = !lean_is_exclusive(x_1); +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; +x_15 = lean_ctor_get(x_1, 0); +x_16 = lean_ctor_get(x_1, 1); +x_17 = lean_ctor_get(x_1, 2); +x_18 = lean_box(0); +lean_inc(x_8); +x_19 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_16, x_8, x_18); +x_20 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_17, x_8); +lean_dec(x_8); +lean_inc(x_15); +x_21 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_15, x_13); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +lean_ctor_set(x_3, 1, x_22); +lean_ctor_set(x_3, 0, x_2); +x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_15, x_13, x_3); +lean_ctor_set(x_1, 2, x_20); +lean_ctor_set(x_1, 1, x_19); +lean_ctor_set(x_1, 0, x_23); +return x_1; } else { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec(x_21); +lean_ctor_set(x_3, 1, x_24); +lean_ctor_set(x_3, 0, x_2); +x_25 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_15, x_13, x_3); +lean_ctor_set(x_1, 2, x_20); +lean_ctor_set(x_1, 1, x_19); +lean_ctor_set(x_1, 0, x_25); +return x_1; } } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_26 = lean_ctor_get(x_1, 0); +x_27 = lean_ctor_get(x_1, 1); +x_28 = lean_ctor_get(x_1, 2); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_1); +x_29 = lean_box(0); +lean_inc(x_8); +x_30 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_27, x_8, x_29); +x_31 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_28, x_8); +lean_dec(x_8); +lean_inc(x_26); +x_32 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_26, x_13); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_box(0); +lean_ctor_set(x_3, 1, x_33); +lean_ctor_set(x_3, 0, x_2); +x_34 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_26, x_13, x_3); +x_35 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_30); +lean_ctor_set(x_35, 2, x_31); +return x_35; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; -} -} -} +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_32, 0); +lean_inc(x_36); +lean_dec(x_32); +lean_ctor_set(x_3, 1, x_36); +lean_ctor_set(x_3, 0, x_2); +x_37 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_26, x_13, x_3); +x_38 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_30); +lean_ctor_set(x_38, 2, x_31); +return x_38; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_array_uget(x_3, x_5); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_39 = lean_ctor_get(x_3, 1); +lean_inc(x_39); +lean_dec(x_3); +x_40 = lean_ctor_get(x_6, 0); +lean_inc(x_40); +lean_dec(x_6); lean_inc(x_8); -x_9 = l_Lean_Meta_Grind_ppPattern(x_8); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_ctor_set(x_2, 4, x_39); +x_41 = lean_ctor_get(x_1, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_1, 1); +lean_inc(x_42); +x_43 = lean_ctor_get(x_1, 2); +lean_inc(x_43); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + x_44 = x_1; +} else { + lean_dec_ref(x_1); + x_44 = lean_box(0); +} +x_45 = lean_box(0); +lean_inc(x_8); +x_46 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_42, x_8, x_45); +x_47 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_43, x_8); lean_dec(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; -x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; -x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); -x_15 = lean_box(0); -x_16 = lean_apply_3(x_10, x_14, x_6, x_15); -if (lean_obj_tag(x_16) == 0) +lean_inc(x_41); +x_48 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_41, x_40); +if (lean_obj_tag(x_48) == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -return x_17; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_49 = lean_box(0); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_2); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_41, x_40, x_50); +if (lean_is_scalar(x_44)) { + x_52 = lean_alloc_ctor(0, 3, 0); +} else { + x_52 = x_44; +} +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_46); +lean_ctor_set(x_52, 2, x_47); +return x_52; } else { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_5, x_19); -x_5 = x_20; -x_6 = x_18; -goto _start; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = lean_ctor_get(x_48, 0); +lean_inc(x_53); +lean_dec(x_48); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_2); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_41, x_40, x_54); +if (lean_is_scalar(x_44)) { + x_56 = lean_alloc_ctor(0, 3, 0); +} else { + x_56 = x_44; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_46); +lean_ctor_set(x_56, 2, x_47); +return x_56; +} } } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_apply_3(x_10, x_9, x_6, x_22); -if (lean_obj_tag(x_23) == 0) +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_57 = lean_ctor_get(x_2, 0); +x_58 = lean_ctor_get(x_2, 1); +x_59 = lean_ctor_get(x_2, 2); +x_60 = lean_ctor_get(x_2, 3); +x_61 = lean_ctor_get(x_2, 5); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_2); +x_62 = lean_ctor_get(x_3, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_3)) { + lean_ctor_release(x_3, 0); + lean_ctor_release(x_3, 1); + x_63 = x_3; +} else { + lean_dec_ref(x_3); + x_63 = lean_box(0); +} +x_64 = lean_ctor_get(x_6, 0); +lean_inc(x_64); +lean_dec(x_6); +lean_inc(x_61); +x_65 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_65, 0, x_57); +lean_ctor_set(x_65, 1, x_58); +lean_ctor_set(x_65, 2, x_59); +lean_ctor_set(x_65, 3, x_60); +lean_ctor_set(x_65, 4, x_62); +lean_ctor_set(x_65, 5, x_61); +x_66 = lean_ctor_get(x_1, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_1, 1); +lean_inc(x_67); +x_68 = lean_ctor_get(x_1, 2); +lean_inc(x_68); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + x_69 = x_1; +} else { + lean_dec_ref(x_1); + x_69 = lean_box(0); +} +x_70 = lean_box(0); +lean_inc(x_61); +x_71 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_67, x_61, x_70); +x_72 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_68, x_61); +lean_dec(x_61); +lean_inc(x_66); +x_73 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_66, x_64); +if (lean_obj_tag(x_73) == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_74 = lean_box(0); +if (lean_is_scalar(x_63)) { + x_75 = lean_alloc_ctor(1, 2, 0); +} else { + x_75 = x_63; +} +lean_ctor_set(x_75, 0, x_65); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_66, x_64, x_75); +if (lean_is_scalar(x_69)) { + x_77 = lean_alloc_ctor(0, 3, 0); +} else { + x_77 = x_69; +} +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_71); +lean_ctor_set(x_77, 2, x_72); +return x_77; } else { -lean_object* x_25; size_t x_26; size_t x_27; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = 1; -x_27 = lean_usize_add(x_5, x_26); -x_5 = x_27; -x_6 = x_25; -goto _start; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_73, 0); +lean_inc(x_78); +lean_dec(x_73); +if (lean_is_scalar(x_63)) { + x_79 = lean_alloc_ctor(1, 2, 0); +} else { + x_79 = x_63; } +lean_ctor_set(x_79, 0, x_65); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__12(x_66, x_64, x_79); +if (lean_is_scalar(x_69)) { + x_81 = lean_alloc_ctor(0, 3, 0); +} else { + x_81 = x_69; } +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_71); +lean_ctor_set(x_81, 2, x_72); +return x_81; } } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("#", 1, 1); -return x_1; +lean_object* x_82; lean_object* x_83; +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_82 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4; +x_83 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1(x_82); +return x_83; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__3() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__4(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__4() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3(x_1, x_6, x_7, x_4, x_5); +return x_8; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__5() { +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_4; +x_4 = l_Array_indexOfAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__8(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__6() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\?", 1, 1); -return x_1; +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__7(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__7() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__6; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_1, x_2); +lean_dec(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__8() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`[", 2, 2); -return x_1; +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__11(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_6; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__9() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__8; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__10(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__10() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("]", 1, 1); -return x_1; +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_1, x_2); +lean_dec(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__11() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_ppPattern___closed__10; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__14(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); -if (lean_obj_tag(x_2) == 0) +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__13(x_1, x_6, x_7, x_4, x_5); +return x_8; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_4 = lean_expr_eqv(x_1, x_3); -if (x_4 == 0) +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) { -switch (lean_obj_tag(x_1)) { -case 0: +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; +} +else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_Origin_key(x_5); +x_11 = l_Lean_Meta_Grind_Origin_key(x_9); +lean_dec(x_9); +x_12 = lean_name_eq(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_4, x_13); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_14; +goto _start; +} +else +{ +uint8_t x_16; +lean_dec(x_4); +x_16 = 1; +return x_16; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); lean_dec(x_1); -x_6 = l___private_Init_Data_Repr_0__Nat_reprFast(x_5); -x_7 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_7, 0, x_6); -x_8 = l_Lean_MessageData_ofFormat(x_7); -x_9 = l_Lean_Meta_Grind_ppPattern___closed__2; -x_10 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -x_11 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_12 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_Origin_key(x_3); +x_13 = l_Lean_Meta_Grind_Origin_key(x_11); +lean_dec(x_11); +x_14 = lean_name_eq(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +return x_14; } case 1: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; -x_13 = l_Lean_Expr_getAppFn(x_1); -x_14 = l_Lean_MessageData_ofExpr(x_13); -x_15 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_16 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -x_17 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -x_18 = lean_unsigned_to_nat(0u); -x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_18); -x_20 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_object* x_15; size_t x_16; +x_15 = lean_ctor_get(x_10, 0); +lean_inc(x_15); +lean_dec(x_10); +x_16 = lean_usize_shift_right(x_2, x_5); +x_1 = x_15; +x_2 = x_16; +goto _start; +} +default: +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_1, 0); lean_inc(x_19); -x_21 = lean_mk_array(x_19, x_20); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_sub(x_19, x_22); +x_20 = lean_ctor_get(x_1, 1); +lean_inc(x_20); +lean_dec(x_1); +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3(x_19, x_20, lean_box(0), x_21, x_3); +lean_dec(x_20); lean_dec(x_19); -x_24 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_21, x_23); -x_25 = lean_box(0); -x_26 = lean_array_size(x_24); -x_27 = 0; -x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_24, x_25, x_24, x_26, x_27, x_17); -lean_dec(x_24); -return x_28; +return x_22; } -case 2: +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; size_t x_43; lean_object* x_44; -x_29 = l_Lean_Expr_getAppFn(x_1); -x_30 = l_Lean_MessageData_ofExpr(x_29); -x_31 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = lean_unsigned_to_nat(0u); -x_35 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_34); -x_36 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_35); -x_37 = lean_mk_array(x_35, x_36); -x_38 = lean_unsigned_to_nat(1u); -x_39 = lean_nat_sub(x_35, x_38); -lean_dec(x_35); -x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_37, x_39); -x_41 = lean_box(0); -x_42 = lean_array_size(x_40); -x_43 = 0; -x_44 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_40, x_41, x_40, x_42, x_43, x_33); -lean_dec(x_40); -return x_44; +lean_object* x_3; uint64_t x_4; size_t x_5; uint8_t x_6; +x_3 = l_Lean_Meta_Grind_Origin_key(x_2); +x_4 = l_Lean_Name_hash___override(x_3); +lean_dec(x_3); +x_5 = lean_uint64_to_usize(x_4); +x_6 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2(x_1, x_5, x_2); +return x_6; } -case 3: +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_EMatchTheorems_contains(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; -x_45 = l_Lean_Expr_getAppFn(x_1); -x_46 = l_Lean_MessageData_ofExpr(x_45); -x_47 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -x_50 = lean_unsigned_to_nat(0u); -x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); -x_52 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_51); -x_53 = lean_mk_array(x_51, x_52); -x_54 = lean_unsigned_to_nat(1u); -x_55 = lean_nat_sub(x_51, x_54); -lean_dec(x_51); -x_56 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_53, x_55); -x_57 = lean_box(0); -x_58 = lean_array_size(x_56); -x_59 = 0; -x_60 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_56, x_57, x_56, x_58, x_59, x_49); -lean_dec(x_56); -return x_60; +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(x_3, x_2); +return x_4; } -case 4: +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; -x_61 = l_Lean_Expr_getAppFn(x_1); -x_62 = l_Lean_MessageData_ofExpr(x_61); -x_63 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_62); -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_63); -x_66 = lean_unsigned_to_nat(0u); -x_67 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_66); -x_68 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_67); -x_69 = lean_mk_array(x_67, x_68); -x_70 = lean_unsigned_to_nat(1u); -x_71 = lean_nat_sub(x_67, x_70); -lean_dec(x_67); -x_72 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_69, x_71); -x_73 = lean_box(0); -x_74 = lean_array_size(x_72); -x_75 = 0; -x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_72, x_73, x_72, x_74, x_75, x_65); -lean_dec(x_72); -return x_76; +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; } -case 5: -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; size_t x_90; size_t x_91; lean_object* x_92; -x_77 = l_Lean_Expr_getAppFn(x_1); -x_78 = l_Lean_MessageData_ofExpr(x_77); -x_79 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_78); -x_81 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_79); -x_82 = lean_unsigned_to_nat(0u); -x_83 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_82); -x_84 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_83); -x_85 = lean_mk_array(x_83, x_84); -x_86 = lean_unsigned_to_nat(1u); -x_87 = lean_nat_sub(x_83, x_86); -lean_dec(x_83); -x_88 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_85, x_87); -x_89 = lean_box(0); -x_90 = lean_array_size(x_88); -x_91 = 0; -x_92 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_88, x_89, x_88, x_90, x_91, x_81); -lean_dec(x_88); -return x_92; } -case 6: +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; size_t x_106; size_t x_107; lean_object* x_108; -x_93 = l_Lean_Expr_getAppFn(x_1); -x_94 = l_Lean_MessageData_ofExpr(x_93); -x_95 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_96 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_94); -x_97 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_95); -x_98 = lean_unsigned_to_nat(0u); -x_99 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_98); -x_100 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_99); -x_101 = lean_mk_array(x_99, x_100); -x_102 = lean_unsigned_to_nat(1u); -x_103 = lean_nat_sub(x_99, x_102); -lean_dec(x_99); -x_104 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_101, x_103); -x_105 = lean_box(0); -x_106 = lean_array_size(x_104); -x_107 = 0; -x_108 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_104, x_105, x_104, x_106, x_107, x_97); -lean_dec(x_104); -return x_108; +size_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; } -case 7: -{ -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; size_t x_122; size_t x_123; lean_object* x_124; -x_109 = l_Lean_Expr_getAppFn(x_1); -x_110 = l_Lean_MessageData_ofExpr(x_109); -x_111 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_112 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_110); -x_113 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = lean_unsigned_to_nat(0u); -x_115 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_114); -x_116 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_115); -x_117 = lean_mk_array(x_115, x_116); -x_118 = lean_unsigned_to_nat(1u); -x_119 = lean_nat_sub(x_115, x_118); -lean_dec(x_115); -x_120 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_117, x_119); -x_121 = lean_box(0); -x_122 = lean_array_size(x_120); -x_123 = 0; -x_124 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_120, x_121, x_120, x_122, x_123, x_113); -lean_dec(x_120); -return x_124; } -case 8: +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; size_t x_138; size_t x_139; lean_object* x_140; -x_125 = l_Lean_Expr_getAppFn(x_1); -x_126 = l_Lean_MessageData_ofExpr(x_125); -x_127 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_128 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_126); -x_129 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_127); -x_130 = lean_unsigned_to_nat(0u); -x_131 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_130); -x_132 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_131); -x_133 = lean_mk_array(x_131, x_132); -x_134 = lean_unsigned_to_nat(1u); -x_135 = lean_nat_sub(x_131, x_134); -lean_dec(x_131); -x_136 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_133, x_135); -x_137 = lean_box(0); -x_138 = lean_array_size(x_136); -x_139 = 0; -x_140 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_136, x_137, x_136, x_138, x_139, x_129); -lean_dec(x_136); -return x_140; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } -case 9: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_contains___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; size_t x_154; size_t x_155; lean_object* x_156; -x_141 = l_Lean_Expr_getAppFn(x_1); -x_142 = l_Lean_MessageData_ofExpr(x_141); -x_143 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_144 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_142); -x_145 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_145, 0, x_144); -lean_ctor_set(x_145, 1, x_143); -x_146 = lean_unsigned_to_nat(0u); -x_147 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_146); -x_148 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_147); -x_149 = lean_mk_array(x_147, x_148); -x_150 = lean_unsigned_to_nat(1u); -x_151 = lean_nat_sub(x_147, x_150); -lean_dec(x_147); -x_152 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_149, x_151); -x_153 = lean_box(0); -x_154 = lean_array_size(x_152); -x_155 = 0; -x_156 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_152, x_153, x_152, x_154, x_155, x_145); -lean_dec(x_152); -return x_156; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_EMatchTheorems_contains(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } -case 10: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_erase(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; size_t x_170; size_t x_171; lean_object* x_172; -x_157 = l_Lean_Expr_getAppFn(x_1); -x_158 = l_Lean_MessageData_ofExpr(x_157); -x_159 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_160 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_158); -x_161 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_159); -x_162 = lean_unsigned_to_nat(0u); -x_163 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_162); -x_164 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_163); -x_165 = lean_mk_array(x_163, x_164); -x_166 = lean_unsigned_to_nat(1u); -x_167 = lean_nat_sub(x_163, x_166); -lean_dec(x_163); -x_168 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_165, x_167); -x_169 = lean_box(0); -x_170 = lean_array_size(x_168); -x_171 = 0; -x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_168, x_169, x_168, x_170, x_171, x_161); -lean_dec(x_168); -return x_172; +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_4, x_2); +x_7 = lean_box(0); +x_8 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_5, x_2, x_7); +lean_ctor_set(x_1, 2, x_8); +lean_ctor_set(x_1, 1, x_6); +return x_1; } -default: +else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; size_t x_186; size_t x_187; lean_object* x_188; -x_173 = l_Lean_Expr_getAppFn(x_1); -x_174 = l_Lean_MessageData_ofExpr(x_173); -x_175 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_176 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_176, 0, x_175); -lean_ctor_set(x_176, 1, x_174); -x_177 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_177, 0, x_176); -lean_ctor_set(x_177, 1, x_175); -x_178 = lean_unsigned_to_nat(0u); -x_179 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_178); -x_180 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_179); -x_181 = lean_mk_array(x_179, x_180); -x_182 = lean_unsigned_to_nat(1u); -x_183 = lean_nat_sub(x_179, x_182); -lean_dec(x_179); -x_184 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_181, x_183); -x_185 = lean_box(0); -x_186 = lean_array_size(x_184); -x_187 = 0; -x_188 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_184, x_185, x_184, x_186, x_187, x_177); -lean_dec(x_184); -return x_188; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_ctor_get(x_1, 2); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_12 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6(x_10, x_2); +x_13 = lean_box(0); +x_14 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_11, x_2, x_13); +x_15 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_15, 0, x_9); +lean_ctor_set(x_15, 1, x_12); +lean_ctor_set(x_15, 2, x_14); +return x_15; } } } -else +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_EMatchTheorems_isErased(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_189; +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_1, 2); +lean_inc(x_3); lean_dec(x_1); -x_189 = l_Lean_Meta_Grind_ppPattern___closed__7; -return x_189; +x_4 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_EMatchTheorems_contains___spec__1(x_3, x_2); +return x_4; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_isErased___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -lean_dec(x_1); -x_190 = lean_ctor_get(x_2, 0); -lean_inc(x_190); +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_EMatchTheorems_isErased(x_1, x_2); lean_dec(x_2); -x_191 = l_Lean_MessageData_ofExpr(x_190); -x_192 = l_Lean_Meta_Grind_ppPattern___closed__9; -x_193 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_193, 0, x_192); -lean_ctor_set(x_193, 1, x_191); -x_194 = l_Lean_Meta_Grind_ppPattern___closed__11; -x_195 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -return x_195; -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); +x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +if (lean_obj_tag(x_1) == 0) { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +switch (lean_obj_tag(x_10)) { +case 0: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_name_eq(x_3, x_11); +lean_dec(x_11); +if (x_12 == 0) { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); +lean_dec(x_8); lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} +return x_1; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +uint8_t x_13; +x_13 = !lean_is_exclusive(x_1); +if (x_13 == 0) { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_1, 0); +lean_dec(x_14); +x_15 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_15); +return x_1; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_16; lean_object* x_17; lean_dec(x_1); -return x_9; -} +x_16 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +case 1: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +uint8_t x_18; +x_18 = !lean_is_exclusive(x_1); +if (x_18 == 0) { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object* x_1, lean_object* x_2) { -_start: +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_1, 0); +lean_dec(x_19); +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) { -if (lean_obj_tag(x_2) == 0) +lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_array_set(x_4, x_8, x_9); +x_23 = lean_usize_shift_right(x_2, x_5); +x_24 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_21, x_23, x_3); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); +if (lean_obj_tag(x_25) == 0) { -uint8_t x_3; -x_3 = 0; -return x_3; +lean_object* x_26; +lean_ctor_set(x_10, 0, x_24); +x_26 = lean_array_set(x_22, x_8, x_10); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_26); +return x_1; } else { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_4, x_1); -if (x_6 == 0) +lean_object* x_27; uint8_t x_28; +lean_dec(x_24); +lean_free_object(x_10); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -x_2 = x_5; -goto _start; +lean_object* x_29; +x_29 = lean_array_set(x_22, x_8, x_27); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_29); +return x_1; } else { -uint8_t x_8; -x_8 = 1; -return x_8; -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_27, 0); +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_27); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_array_set(x_22, x_8, x_32); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_33); +return x_1; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_2) == 0) +lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_10, 0); +lean_inc(x_34); +lean_dec(x_10); +x_35 = lean_array_set(x_4, x_8, x_9); +x_36 = lean_usize_shift_right(x_2, x_5); +x_37 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_34, x_36, x_3); +lean_inc(x_37); +x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); +if (lean_obj_tag(x_38) == 0) { +lean_object* x_39; lean_object* x_40; +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_37); +x_40 = lean_array_set(x_35, x_8, x_39); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_40); return x_1; } else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = l_Lean_HeadIndex_hash(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_37); +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(x_38); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_44 = x_41; +} else { + lean_dec_ref(x_41); + x_44 = lean_box(0); +} +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(0, 2, 0); +} else { + x_45 = x_44; +} +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +x_46 = lean_array_set(x_35, x_8, x_45); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_46); +return x_1; +} +} } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = l_Lean_HeadIndex_hash(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; +lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_1); +x_47 = lean_ctor_get(x_10, 0); +lean_inc(x_47); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_48 = x_10; +} else { + lean_dec_ref(x_10); + x_48 = lean_box(0); +} +x_49 = lean_array_set(x_4, x_8, x_9); +x_50 = lean_usize_shift_right(x_2, x_5); +x_51 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_47, x_50, x_3); +lean_inc(x_51); +x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +if (lean_is_scalar(x_48)) { + x_53 = lean_alloc_ctor(1, 1, 0); +} else { + x_53 = x_48; +} +lean_ctor_set(x_53, 0, x_51); +x_54 = lean_array_set(x_49, x_8, x_53); +lean_dec(x_8); +x_55 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_55, 0, x_54); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_51); +lean_dec(x_48); +x_56 = lean_ctor_get(x_52, 0); +lean_inc(x_56); +lean_dec(x_52); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; +} else { + lean_dec_ref(x_56); + x_59 = lean_box(0); +} +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 2, 0); +} else { + x_60 = x_59; } +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_58); +x_61 = lean_array_set(x_49, x_8, x_60); +lean_dec(x_8); +x_62 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_62, 0, x_61); +return x_62; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +default: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_8); lean_dec(x_4); -if (x_5 == 0) +return x_1; +} +} +} +else { -lean_dec(x_2); -lean_dec(x_1); -return x_3; +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_1, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 1); +lean_inc(x_64); +x_65 = lean_unsigned_to_nat(0u); +x_66 = l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(x_63, x_3, x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_dec(x_64); +lean_dec(x_63); +return x_1; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); +uint8_t x_67; +x_67 = !lean_is_exclusive(x_1); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_68 = lean_ctor_get(x_1, 1); +lean_dec(x_68); +x_69 = lean_ctor_get(x_1, 0); +lean_dec(x_69); +x_70 = lean_ctor_get(x_66, 0); +lean_inc(x_70); +lean_dec(x_66); +lean_inc(x_70); +x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); +x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); +lean_ctor_set(x_1, 1, x_72); +lean_ctor_set(x_1, 0, x_71); +return x_1; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; +x_73 = lean_ctor_get(x_66, 0); +lean_inc(x_73); +lean_dec(x_66); +lean_inc(x_73); +x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); +x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(lean_object* x_1) { +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(x_7, x_1, x_6); -return x_8; +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Name_hash___override(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_1, x_4, x_2); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_st_ref_get(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; -x_12 = lean_ctor_get(x_8, 1); -x_13 = lean_ctor_get(x_8, 0); -lean_dec(x_13); -x_14 = lean_ctor_get(x_10, 1); -lean_inc(x_14); -lean_dec(x_10); -x_15 = lean_array_get_size(x_14); -x_16 = l_Lean_HeadIndex_hash(x_1); -x_17 = 32; -x_18 = lean_uint64_shift_right(x_16, x_17); -x_19 = lean_uint64_xor(x_16, x_18); -x_20 = 16; -x_21 = lean_uint64_shift_right(x_19, x_20); -x_22 = lean_uint64_xor(x_19, x_21); -x_23 = lean_uint64_to_usize(x_22); -x_24 = lean_usize_of_nat(x_15); -lean_dec(x_15); -x_25 = 1; -x_26 = lean_usize_sub(x_24, x_25); -x_27 = lean_usize_land(x_23, x_26); -x_28 = lean_array_uget(x_14, x_27); -lean_dec(x_14); -x_29 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_28); -lean_dec(x_28); -if (x_29 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_7 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_4, x_2); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_free_object(x_8); -x_30 = lean_st_ref_take(x_2, x_12); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = !lean_is_exclusive(x_31); -if (x_33 == 0) +lean_object* x_8; +lean_free_object(x_1); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; +} +else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_34 = lean_ctor_get(x_31, 0); -x_35 = lean_ctor_get(x_31, 1); -lean_inc(x_1); -x_36 = lean_array_push(x_34, x_1); -x_37 = !lean_is_exclusive(x_35); -if (x_37 == 0) +uint8_t x_9; +x_9 = !lean_is_exclusive(x_7); +if (x_9 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; uint8_t x_45; -x_38 = lean_ctor_get(x_35, 0); -x_39 = lean_ctor_get(x_35, 1); -x_40 = lean_array_get_size(x_39); -x_41 = lean_usize_of_nat(x_40); -lean_dec(x_40); -x_42 = lean_usize_sub(x_41, x_25); -x_43 = lean_usize_land(x_23, x_42); -x_44 = lean_array_uget(x_39, x_43); -x_45 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_44); -if (x_45 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_7, 0); +x_11 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_4, x_2); +lean_ctor_set(x_1, 0, x_11); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_1); +lean_ctor_set(x_7, 0, x_12); +return x_7; +} +else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_38, x_46); -lean_dec(x_38); -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_1); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_44); -x_50 = lean_array_uset(x_39, x_43, x_49); -x_51 = lean_unsigned_to_nat(4u); -x_52 = lean_nat_mul(x_47, x_51); -x_53 = lean_unsigned_to_nat(3u); -x_54 = lean_nat_div(x_52, x_53); -lean_dec(x_52); -x_55 = lean_array_get_size(x_50); -x_56 = lean_nat_dec_le(x_54, x_55); -lean_dec(x_55); -lean_dec(x_54); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_57 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_50); -lean_ctor_set(x_35, 1, x_57); -lean_ctor_set(x_35, 0, x_47); -lean_ctor_set(x_31, 0, x_36); -x_58 = lean_st_ref_set(x_2, x_31, x_32); -x_59 = !lean_is_exclusive(x_58); -if (x_59 == 0) -{ -lean_object* x_60; -x_60 = lean_ctor_get(x_58, 0); -lean_dec(x_60); -lean_ctor_set(x_58, 0, x_48); -return x_58; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_7, 0); +lean_inc(x_13); +lean_dec(x_7); +x_14 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_4, x_2); +lean_ctor_set(x_1, 0, x_14); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_1); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +return x_16; } -else -{ -lean_object* x_61; lean_object* x_62; -x_61 = lean_ctor_get(x_58, 1); -lean_inc(x_61); -lean_dec(x_58); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_48); -lean_ctor_set(x_62, 1, x_61); -return x_62; } } else { -lean_object* x_63; uint8_t x_64; -lean_ctor_set(x_35, 1, x_50); -lean_ctor_set(x_35, 0, x_47); -lean_ctor_set(x_31, 0, x_36); -x_63 = lean_st_ref_set(x_2, x_31, x_32); -x_64 = !lean_is_exclusive(x_63); -if (x_64 == 0) +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_ctor_get(x_1, 1); +x_19 = lean_ctor_get(x_1, 2); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_1); +lean_inc(x_17); +x_20 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_17, x_2); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_65; -x_65 = lean_ctor_get(x_63, 0); -lean_dec(x_65); -lean_ctor_set(x_63, 0, x_48); -return x_63; +lean_object* x_21; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +x_21 = lean_box(0); +return x_21; } else { -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_63, 1); -lean_inc(x_66); -lean_dec(x_63); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_48); -lean_ctor_set(x_67, 1, x_66); -return x_67; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + x_23 = x_20; +} else { + lean_dec_ref(x_20); + x_23 = lean_box(0); } +x_24 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_17, x_2); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_18); +lean_ctor_set(x_25, 2, x_19); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_22); +lean_ctor_set(x_26, 1, x_25); +if (lean_is_scalar(x_23)) { + x_27 = lean_alloc_ctor(1, 1, 0); +} else { + x_27 = x_23; } +lean_ctor_set(x_27, 0, x_26); +return x_27; } -else -{ -lean_object* x_68; uint8_t x_69; -lean_dec(x_44); -lean_dec(x_1); -lean_ctor_set(x_31, 0, x_36); -x_68 = lean_st_ref_set(x_2, x_31, x_32); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_68, 0); -lean_dec(x_70); -x_71 = lean_box(0); -lean_ctor_set(x_68, 0, x_71); -return x_68; +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_68, 1); -lean_inc(x_72); -lean_dec(x_68); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_72); -return x_74; +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f(x_1, x_2); +lean_dec(x_2); +return x_3; +} } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; size_t x_80; lean_object* x_81; uint8_t x_82; -x_75 = lean_ctor_get(x_35, 0); -x_76 = lean_ctor_get(x_35, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_35); -x_77 = lean_array_get_size(x_76); -x_78 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_79 = lean_usize_sub(x_78, x_25); -x_80 = lean_usize_land(x_23, x_79); -x_81 = lean_array_uget(x_76, x_80); -x_82 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_81); -if (x_82 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +x_8 = l_Lean_Expr_isConst(x_7); +if (x_8 == 0) { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_83 = lean_unsigned_to_nat(1u); -x_84 = lean_nat_add(x_75, x_83); -lean_dec(x_75); -x_85 = lean_box(0); -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_1); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set(x_86, 2, x_81); -x_87 = lean_array_uset(x_76, x_80, x_86); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_mul(x_84, x_88); -x_90 = lean_unsigned_to_nat(3u); -x_91 = lean_nat_div(x_89, x_90); -lean_dec(x_89); -x_92 = lean_array_get_size(x_87); -x_93 = lean_nat_dec_le(x_91, x_92); -lean_dec(x_92); -lean_dec(x_91); -if (x_93 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = l_Array_isEmpty___rarg(x_9); +if (x_10 == 0) { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_94 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_87); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_84); -lean_ctor_set(x_95, 1, x_94); -lean_ctor_set(x_31, 1, x_95); -lean_ctor_set(x_31, 0, x_36); -x_96 = lean_st_ref_set(x_2, x_31, x_32); -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_98 = x_96; -} else { - lean_dec_ref(x_96); - x_98 = lean_box(0); -} -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_98; -} -lean_ctor_set(x_99, 0, x_85); -lean_ctor_set(x_99, 1, x_97); -return x_99; +size_t x_11; size_t x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_array_size(x_9); +x_12 = 0; +lean_inc(x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_11, x_12, x_9, x_2, x_3, x_4, x_5, x_6); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_15); +lean_dec(x_7); +lean_ctor_set(x_13, 0, x_16); +return x_13; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_84); -lean_ctor_set(x_100, 1, x_87); -lean_ctor_set(x_31, 1, x_100); -lean_ctor_set(x_31, 0, x_36); -x_101 = lean_st_ref_set(x_2, x_31, x_32); -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_103 = x_101; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_9, x_17); +lean_dec(x_7); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(0, 2, 0); -} else { - x_104 = x_103; } -lean_ctor_set(x_104, 0, x_85); -lean_ctor_set(x_104, 1, x_102); -return x_104; +else +{ +lean_object* x_21; +lean_dec(x_9); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_7); +lean_ctor_set(x_21, 1, x_6); +return x_21; } } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -lean_dec(x_81); +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); lean_dec(x_1); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_75); -lean_ctor_set(x_105, 1, x_76); -lean_ctor_set(x_31, 1, x_105); -lean_ctor_set(x_31, 0, x_36); -x_106 = lean_st_ref_set(x_2, x_31, x_32); -x_107 = lean_ctor_get(x_106, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_108 = x_106; -} else { - lean_dec_ref(x_106); - x_108 = lean_box(0); -} -x_109 = lean_box(0); -if (lean_is_scalar(x_108)) { - x_110 = lean_alloc_ctor(0, 2, 0); -} else { - x_110 = x_108; +x_23 = l_Array_isEmpty___rarg(x_22); +if (x_23 == 0) +{ +size_t x_24; size_t x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_array_size(x_22); +x_25 = 0; +lean_inc(x_22); +x_26 = l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(x_24, x_25, x_22, x_2, x_3, x_4, x_5, x_6); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 0); +x_29 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_28); +lean_dec(x_7); +lean_ctor_set(x_26, 0, x_29); +return x_26; } -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_107); -return x_110; +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_26, 0); +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_26); +x_32 = l_Lean_Expr_instantiateLevelParamsArray(x_7, x_22, x_30); +lean_dec(x_7); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } +else +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_22); +x_34 = l_Lean_Expr_constName_x21(x_7); +lean_inc(x_34); +x_35 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_34, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_35) == 0) +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_37 = lean_ctor_get(x_35, 0); +x_38 = lean_ctor_get(x_35, 1); +x_39 = l_Lean_ConstantInfo_levelParams(x_37); +lean_dec(x_37); +x_40 = l_List_isEmpty___rarg(x_39); +lean_dec(x_39); +if (x_40 == 0) +{ +lean_object* x_41; +lean_free_object(x_35); +lean_dec(x_7); +x_41 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_38); +return x_41; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; size_t x_119; size_t x_120; size_t x_121; lean_object* x_122; uint8_t x_123; -x_111 = lean_ctor_get(x_31, 0); -x_112 = lean_ctor_get(x_31, 1); -x_113 = lean_ctor_get(x_31, 2); -lean_inc(x_113); -lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_31); -lean_inc(x_1); -x_114 = lean_array_push(x_111, x_1); -x_115 = lean_ctor_get(x_112, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_112, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_117 = x_112; -} else { - lean_dec_ref(x_112); - x_117 = lean_box(0); +lean_dec(x_34); +lean_ctor_set(x_35, 0, x_7); +return x_35; } -x_118 = lean_array_get_size(x_116); -x_119 = lean_usize_of_nat(x_118); -lean_dec(x_118); -x_120 = lean_usize_sub(x_119, x_25); -x_121 = lean_usize_land(x_23, x_120); -x_122 = lean_array_uget(x_116, x_121); -x_123 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_122); -if (x_123 == 0) +} +else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_124 = lean_unsigned_to_nat(1u); -x_125 = lean_nat_add(x_115, x_124); -lean_dec(x_115); -x_126 = lean_box(0); -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_1); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_122); -x_128 = lean_array_uset(x_116, x_121, x_127); -x_129 = lean_unsigned_to_nat(4u); -x_130 = lean_nat_mul(x_125, x_129); -x_131 = lean_unsigned_to_nat(3u); -x_132 = lean_nat_div(x_130, x_131); -lean_dec(x_130); -x_133 = lean_array_get_size(x_128); -x_134 = lean_nat_dec_le(x_132, x_133); -lean_dec(x_133); -lean_dec(x_132); -if (x_134 == 0) +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_42 = lean_ctor_get(x_35, 0); +x_43 = lean_ctor_get(x_35, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_35); +x_44 = l_Lean_ConstantInfo_levelParams(x_42); +lean_dec(x_42); +x_45 = l_List_isEmpty___rarg(x_44); +lean_dec(x_44); +if (x_45 == 0) { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_135 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_128); -if (lean_is_scalar(x_117)) { - x_136 = lean_alloc_ctor(0, 2, 0); -} else { - x_136 = x_117; +lean_object* x_46; +lean_dec(x_7); +x_46 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_34, x_2, x_3, x_4, x_5, x_43); +return x_46; } -lean_ctor_set(x_136, 0, x_125); -lean_ctor_set(x_136, 1, x_135); -x_137 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_137, 0, x_114); -lean_ctor_set(x_137, 1, x_136); -lean_ctor_set(x_137, 2, x_113); -x_138 = lean_st_ref_set(x_2, x_137, x_32); -x_139 = lean_ctor_get(x_138, 1); -lean_inc(x_139); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - lean_ctor_release(x_138, 1); - x_140 = x_138; -} else { - lean_dec_ref(x_138); - x_140 = lean_box(0); +else +{ +lean_object* x_47; +lean_dec(x_34); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_7); +lean_ctor_set(x_47, 1, x_43); +return x_47; } -if (lean_is_scalar(x_140)) { - x_141 = lean_alloc_ctor(0, 2, 0); -} else { - x_141 = x_140; } -lean_ctor_set(x_141, 0, x_126); -lean_ctor_set(x_141, 1, x_139); -return x_141; } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -if (lean_is_scalar(x_117)) { - x_142 = lean_alloc_ctor(0, 2, 0); -} else { - x_142 = x_117; +uint8_t x_48; +lean_dec(x_34); +lean_dec(x_7); +x_48 = !lean_is_exclusive(x_35); +if (x_48 == 0) +{ +return x_35; } -lean_ctor_set(x_142, 0, x_125); -lean_ctor_set(x_142, 1, x_128); -x_143 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_143, 0, x_114); -lean_ctor_set(x_143, 1, x_142); -lean_ctor_set(x_143, 2, x_113); -x_144 = lean_st_ref_set(x_2, x_143, x_32); -x_145 = lean_ctor_get(x_144, 1); -lean_inc(x_145); -if (lean_is_exclusive(x_144)) { - lean_ctor_release(x_144, 0); - lean_ctor_release(x_144, 1); - x_146 = x_144; -} else { - lean_dec_ref(x_144); - x_146 = lean_box(0); +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_35, 0); +x_50 = lean_ctor_get(x_35, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_35); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } -if (lean_is_scalar(x_146)) { - x_147 = lean_alloc_ctor(0, 2, 0); -} else { - x_147 = x_146; } -lean_ctor_set(x_147, 0, x_126); -lean_ctor_set(x_147, 1, x_145); -return x_147; } } -else -{ -lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_122); -lean_dec(x_1); -if (lean_is_scalar(x_117)) { - x_148 = lean_alloc_ctor(0, 2, 0); -} else { - x_148 = x_117; } -lean_ctor_set(x_148, 0, x_115); -lean_ctor_set(x_148, 1, x_116); -x_149 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_149, 0, x_114); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set(x_149, 2, x_113); -x_150 = lean_st_ref_set(x_2, x_149, x_32); -x_151 = lean_ctor_get(x_150, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_150)) { - lean_ctor_release(x_150, 0); - lean_ctor_release(x_150, 1); - x_152 = x_150; -} else { - lean_dec_ref(x_150); - x_152 = lean_box(0); } -x_153 = lean_box(0); -if (lean_is_scalar(x_152)) { - x_154 = lean_alloc_ctor(0, 2, 0); -} else { - x_154 = x_152; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_Grind_EMatchTheorem_getProofWithFreshMVarLevels(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; } -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_151); -return x_154; } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private", 8, 8); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2() { +_start: { -lean_object* x_155; -lean_dec(x_1); -x_155 = lean_box(0); -lean_ctor_set(x_8, 0, x_155); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; uint64_t x_159; uint64_t x_160; uint64_t x_161; uint64_t x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; size_t x_166; size_t x_167; size_t x_168; size_t x_169; size_t x_170; lean_object* x_171; uint8_t x_172; -x_156 = lean_ctor_get(x_8, 1); -lean_inc(x_156); -lean_dec(x_8); -x_157 = lean_ctor_get(x_10, 1); -lean_inc(x_157); -lean_dec(x_10); -x_158 = lean_array_get_size(x_157); -x_159 = l_Lean_HeadIndex_hash(x_1); -x_160 = 32; -x_161 = lean_uint64_shift_right(x_159, x_160); -x_162 = lean_uint64_xor(x_159, x_161); -x_163 = 16; -x_164 = lean_uint64_shift_right(x_162, x_163); -x_165 = lean_uint64_xor(x_162, x_164); -x_166 = lean_uint64_to_usize(x_165); -x_167 = lean_usize_of_nat(x_158); -lean_dec(x_158); -x_168 = 1; -x_169 = lean_usize_sub(x_167, x_168); -x_170 = lean_usize_land(x_166, x_169); -x_171 = lean_array_uget(x_157, x_170); -lean_dec(x_157); -x_172 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_171); -lean_dec(x_171); -if (x_172 == 0) +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3() { +_start: { -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; size_t x_185; size_t x_186; size_t x_187; lean_object* x_188; uint8_t x_189; -x_173 = lean_st_ref_take(x_2, x_156); -x_174 = lean_ctor_get(x_173, 0); -lean_inc(x_174); -x_175 = lean_ctor_get(x_173, 1); -lean_inc(x_175); -lean_dec(x_173); -x_176 = lean_ctor_get(x_174, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_174, 1); -lean_inc(x_177); -x_178 = lean_ctor_get(x_174, 2); -lean_inc(x_178); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - lean_ctor_release(x_174, 2); - x_179 = x_174; -} else { - lean_dec_ref(x_174); - x_179 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_inc(x_1); -x_180 = lean_array_push(x_176, x_1); -x_181 = lean_ctor_get(x_177, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_177, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_177)) { - lean_ctor_release(x_177, 0); - lean_ctor_release(x_177, 1); - x_183 = x_177; -} else { - lean_dec_ref(x_177); - x_183 = lean_box(0); } -x_184 = lean_array_get_size(x_182); -x_185 = lean_usize_of_nat(x_184); -lean_dec(x_184); -x_186 = lean_usize_sub(x_185, x_168); -x_187 = lean_usize_land(x_166, x_186); -x_188 = lean_array_uget(x_182, x_187); -x_189 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_188); -if (x_189 == 0) +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4() { +_start: { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; -x_190 = lean_unsigned_to_nat(1u); -x_191 = lean_nat_add(x_181, x_190); -lean_dec(x_181); -x_192 = lean_box(0); -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_1); -lean_ctor_set(x_193, 1, x_192); -lean_ctor_set(x_193, 2, x_188); -x_194 = lean_array_uset(x_182, x_187, x_193); -x_195 = lean_unsigned_to_nat(4u); -x_196 = lean_nat_mul(x_191, x_195); -x_197 = lean_unsigned_to_nat(3u); -x_198 = lean_nat_div(x_196, x_197); -lean_dec(x_196); -x_199 = lean_array_get_size(x_194); -x_200 = lean_nat_dec_le(x_198, x_199); -lean_dec(x_199); -lean_dec(x_198); -if (x_200 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Meta", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5() { +_start: { -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; -x_201 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_194); -if (lean_is_scalar(x_183)) { - x_202 = lean_alloc_ctor(0, 2, 0); -} else { - x_202 = x_183; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_202, 0, x_191); -lean_ctor_set(x_202, 1, x_201); -if (lean_is_scalar(x_179)) { - x_203 = lean_alloc_ctor(0, 3, 0); -} else { - x_203 = x_179; } -lean_ctor_set(x_203, 0, x_180); -lean_ctor_set(x_203, 1, x_202); -lean_ctor_set(x_203, 2, x_178); -x_204 = lean_st_ref_set(x_2, x_203, x_175); -x_205 = lean_ctor_get(x_204, 1); -lean_inc(x_205); -if (lean_is_exclusive(x_204)) { - lean_ctor_release(x_204, 0); - lean_ctor_release(x_204, 1); - x_206 = x_204; -} else { - lean_dec_ref(x_204); - x_206 = lean_box(0); +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; } -if (lean_is_scalar(x_206)) { - x_207 = lean_alloc_ctor(0, 2, 0); -} else { - x_207 = x_206; } -lean_ctor_set(x_207, 0, x_192); -lean_ctor_set(x_207, 1, x_205); -return x_207; +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8() { +_start: { -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -if (lean_is_scalar(x_183)) { - x_208 = lean_alloc_ctor(0, 2, 0); -} else { - x_208 = x_183; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_208, 0, x_191); -lean_ctor_set(x_208, 1, x_194); -if (lean_is_scalar(x_179)) { - x_209 = lean_alloc_ctor(0, 3, 0); -} else { - x_209 = x_179; } -lean_ctor_set(x_209, 0, x_180); -lean_ctor_set(x_209, 1, x_208); -lean_ctor_set(x_209, 2, x_178); -x_210 = lean_st_ref_set(x_2, x_209, x_175); -x_211 = lean_ctor_get(x_210, 1); -lean_inc(x_211); -if (lean_is_exclusive(x_210)) { - lean_ctor_release(x_210, 0); - lean_ctor_release(x_210, 1); - x_212 = x_210; -} else { - lean_dec_ref(x_210); - x_212 = lean_box(0); +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("EMatchTheorem", 13, 13); +return x_1; } -if (lean_is_scalar(x_212)) { - x_213 = lean_alloc_ctor(0, 2, 0); -} else { - x_213 = x_212; } -lean_ctor_set(x_213, 0, x_192); -lean_ctor_set(x_213, 1, x_211); -return x_213; +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11() { +_start: { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; -lean_dec(x_188); -lean_dec(x_1); -if (lean_is_scalar(x_183)) { - x_214 = lean_alloc_ctor(0, 2, 0); -} else { - x_214 = x_183; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_214, 0, x_181); -lean_ctor_set(x_214, 1, x_182); -if (lean_is_scalar(x_179)) { - x_215 = lean_alloc_ctor(0, 3, 0); -} else { - x_215 = x_179; } -lean_ctor_set(x_215, 0, x_180); -lean_ctor_set(x_215, 1, x_214); -lean_ctor_set(x_215, 2, x_178); -x_216 = lean_st_ref_set(x_2, x_215, x_175); -x_217 = lean_ctor_get(x_216, 1); -lean_inc(x_217); -if (lean_is_exclusive(x_216)) { - lean_ctor_release(x_216, 0); - lean_ctor_release(x_216, 1); - x_218 = x_216; -} else { - lean_dec_ref(x_216); - x_218 = lean_box(0); +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -x_219 = lean_box(0); -if (lean_is_scalar(x_218)) { - x_220 = lean_alloc_ctor(0, 2, 0); -} else { - x_220 = x_218; } -lean_ctor_set(x_220, 0, x_219); -lean_ctor_set(x_220, 1, x_217); -return x_220; +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14() { +_start: { -lean_object* x_221; lean_object* x_222; -lean_dec(x_1); -x_221 = lean_box(0); -x_222 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_222, 0, x_221); -lean_ctor_set(x_222, 1, x_156); -return x_222; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematchTheoremsExt", 17, 17); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16() { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17() { _start: { -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; } } -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18() { _start: { -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_EMatchTheorems_insert), 2, 0); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19() { +_start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_nat_dec_eq(x_4, x_1); -if (x_6 == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20() { +_start: { -x_2 = x_5; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17; +x_4 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666_(lean_object* x_1) { +_start: { -uint8_t x_8; -x_8 = 1; -return x_8; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20; +x_3 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_2, x_1); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2() { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_st_ref_get(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 2); -lean_inc(x_10); -lean_dec(x_9); -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3() { +_start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; -x_12 = lean_ctor_get(x_8, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_array_get_size(x_13); -x_15 = lean_uint64_of_nat(x_1); -x_16 = 32; -x_17 = lean_uint64_shift_right(x_15, x_16); -x_18 = lean_uint64_xor(x_15, x_17); -x_19 = 16; -x_20 = lean_uint64_shift_right(x_18, x_19); -x_21 = lean_uint64_xor(x_18, x_20); -x_22 = lean_uint64_to_usize(x_21); -x_23 = lean_usize_of_nat(x_14); -lean_dec(x_14); -x_24 = 1; -x_25 = lean_usize_sub(x_23, x_24); -x_26 = lean_usize_land(x_22, x_25); -x_27 = lean_array_uget(x_13, x_26); -lean_dec(x_13); -x_28 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_27); -lean_dec(x_27); -x_29 = lean_box(x_28); -lean_ctor_set(x_8, 0, x_29); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("HEq", 3, 3); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4() { +_start: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; uint64_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; size_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; -x_30 = lean_ctor_get(x_8, 1); -lean_inc(x_30); -lean_dec(x_8); -x_31 = lean_ctor_get(x_10, 1); -lean_inc(x_31); -lean_dec(x_10); -x_32 = lean_array_get_size(x_31); -x_33 = lean_uint64_of_nat(x_1); -x_34 = 32; -x_35 = lean_uint64_shift_right(x_33, x_34); -x_36 = lean_uint64_xor(x_33, x_35); -x_37 = 16; -x_38 = lean_uint64_shift_right(x_36, x_37); -x_39 = lean_uint64_xor(x_36, x_38); -x_40 = lean_uint64_to_usize(x_39); -x_41 = lean_usize_of_nat(x_32); -lean_dec(x_32); -x_42 = 1; -x_43 = lean_usize_sub(x_41, x_42); -x_44 = lean_usize_land(x_40, x_43); -x_45 = lean_array_uget(x_31, x_44); -lean_dec(x_31); -x_46 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_45); -lean_dec(x_45); -x_47 = lean_box(x_46); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_30); -return x_48; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Iff", 3, 3); +return x_1; } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6() { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7() { _start: { -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("And", 3, 3); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8() { _start: { -if (lean_obj_tag(x_3) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9() { +_start: { -lean_dec(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Or", 2, 2); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10() { +_start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 2); -x_7 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_5); -x_8 = lean_apply_1(x_1, x_5); -x_9 = lean_unbox_uint64(x_8); -lean_dec(x_8); -x_10 = 32; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = 16; -x_14 = lean_uint64_shift_right(x_12, x_13); -x_15 = lean_uint64_xor(x_12, x_14); -x_16 = lean_uint64_to_usize(x_15); -x_17 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_18 = 1; -x_19 = lean_usize_sub(x_17, x_18); -x_20 = lean_usize_land(x_16, x_19); -x_21 = lean_array_uget(x_2, x_20); -lean_ctor_set(x_3, 2, x_21); -x_22 = lean_array_uset(x_2, x_20, x_3); -x_2 = x_22; -x_3 = x_6; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11() { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_24 = lean_ctor_get(x_3, 0); -x_25 = lean_ctor_get(x_3, 1); -x_26 = lean_ctor_get(x_3, 2); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_3); -x_27 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_24); -x_28 = lean_apply_1(x_1, x_24); -x_29 = lean_unbox_uint64(x_28); -lean_dec(x_28); -x_30 = 32; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = 16; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = lean_uint64_to_usize(x_35); -x_37 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_38 = 1; -x_39 = lean_usize_sub(x_37, x_38); -x_40 = lean_usize_land(x_36, x_39); -x_41 = lean_array_uget(x_2, x_40); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_24); -lean_ctor_set(x_42, 1, x_25); -lean_ctor_set(x_42, 2, x_41); -x_43 = lean_array_uset(x_2, x_40, x_42); -x_2 = x_43; -x_3 = x_26; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Not", 3, 3); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__11; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13() { _start: { -if (lean_obj_tag(x_2) == 0) -{ -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = lean_uint64_of_nat(x_4); -x_8 = 32; -x_9 = lean_uint64_shift_right(x_7, x_8); -x_10 = lean_uint64_xor(x_7, x_9); -x_11 = 16; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = lean_uint64_to_usize(x_13); -x_15 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_16 = 1; -x_17 = lean_usize_sub(x_15, x_16); -x_18 = lean_usize_land(x_14, x_17); -x_19 = lean_array_uget(x_1, x_18); -lean_ctor_set(x_2, 2, x_19); -x_20 = lean_array_uset(x_1, x_18, x_2); -x_1 = x_20; -x_2 = x_5; -goto _start; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14() { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_22 = lean_ctor_get(x_2, 0); -x_23 = lean_ctor_get(x_2, 1); -x_24 = lean_ctor_get(x_2, 2); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_2); -x_25 = lean_array_get_size(x_1); -x_26 = lean_uint64_of_nat(x_22); -x_27 = 32; -x_28 = lean_uint64_shift_right(x_26, x_27); -x_29 = lean_uint64_xor(x_26, x_28); -x_30 = 16; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = lean_uint64_to_usize(x_32); -x_34 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_35 = 1; -x_36 = lean_usize_sub(x_34, x_35); -x_37 = lean_usize_land(x_33, x_36); -x_38 = lean_array_uget(x_1, x_37); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_23); -lean_ctor_set(x_39, 2, x_38); -x_40 = lean_array_uset(x_1, x_37, x_39); -x_1 = x_40; -x_2 = x_24; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__10; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__8; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16() { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17() { +_start: { -lean_dec(x_2); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18() { +_start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__18; +x_2 = lean_array_mk(x_1); +return x_2; +} } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(lean_object* x_1) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames() { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(x_7, x_1, x_6); -return x_8; +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__19; +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(lean_object* x_1) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_8 = lean_st_ref_take(x_2, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 2); -lean_inc(x_10); -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_dec(x_8); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +x_3 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden___boxed(lean_object* x_1) { +_start: { -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_9, 2); -lean_dec(x_13); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isForbidden(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1() { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; uint8_t x_31; -x_15 = lean_ctor_get(x_10, 0); -x_16 = lean_ctor_get(x_10, 1); -x_17 = lean_array_get_size(x_16); -x_18 = lean_uint64_of_nat(x_1); -x_19 = 32; -x_20 = lean_uint64_shift_right(x_18, x_19); -x_21 = lean_uint64_xor(x_18, x_20); -x_22 = 16; -x_23 = lean_uint64_shift_right(x_21, x_22); -x_24 = lean_uint64_xor(x_21, x_23); -x_25 = lean_uint64_to_usize(x_24); -x_26 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_27 = 1; -x_28 = lean_usize_sub(x_26, x_27); -x_29 = lean_usize_land(x_25, x_28); -x_30 = lean_array_uget(x_16, x_29); -x_31 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_30); -if (x_31 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind_dontcare]", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_32 = lean_unsigned_to_nat(1u); -x_33 = lean_nat_add(x_15, x_32); -lean_dec(x_15); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_1); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_30); -x_36 = lean_array_uset(x_16, x_29, x_35); -x_37 = lean_unsigned_to_nat(4u); -x_38 = lean_nat_mul(x_33, x_37); -x_39 = lean_unsigned_to_nat(3u); -x_40 = lean_nat_div(x_38, x_39); -lean_dec(x_38); -x_41 = lean_array_get_size(x_36); -x_42 = lean_nat_dec_le(x_40, x_41); -lean_dec(x_41); -lean_dec(x_40); -if (x_42 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3() { +_start: { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_36); -lean_ctor_set(x_10, 1, x_43); -lean_ctor_set(x_10, 0, x_33); -x_44 = lean_st_ref_set(x_2, x_9, x_11); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare() { +_start: { -lean_object* x_46; -x_46 = lean_ctor_get(x_44, 0); -lean_dec(x_46); -lean_ctor_set(x_44, 0, x_34); -return x_44; +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare___closed__3; +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__1() { +_start: { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_34); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__2() { +_start: { -lean_object* x_49; uint8_t x_50; -lean_ctor_set(x_10, 1, x_36); -lean_ctor_set(x_10, 0, x_33); -x_49 = lean_st_ref_set(x_2, x_9, x_11); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ground_pat", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkGroundPattern___closed__3() { +_start: { -lean_object* x_51; -x_51 = lean_ctor_get(x_49, 0); -lean_dec(x_51); -lean_ctor_set(x_49, 0, x_34); -return x_49; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object* x_1) { +_start: { -lean_object* x_52; lean_object* x_53; -x_52 = lean_ctor_get(x_49, 1); -lean_inc(x_52); -lean_dec(x_49); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_34); -lean_ctor_set(x_53, 1, x_52); -return x_53; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; +x_3 = l_Lean_mkAnnotation(x_2, x_1); +return x_3; +} } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__3; +x_3 = l_Lean_annotation_x3f(x_2, x_1); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_groundPattern_x3f___boxed(lean_object* x_1) { +_start: { -lean_object* x_54; uint8_t x_55; -lean_dec(x_30); +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); lean_dec(x_1); -x_54 = lean_st_ref_set(x_2, x_9, x_11); -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(lean_object* x_1) { +_start: { -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_54, 0); -lean_dec(x_56); -x_57 = lean_box(0); -lean_ctor_set(x_54, 0, x_57); -return x_54; +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_54, 1); -lean_inc(x_58); -lean_dec(x_54); -x_59 = lean_box(0); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_58); -return x_60; +uint8_t x_4; +lean_dec(x_2); +x_4 = 1; +return x_4; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern___boxed(lean_object* x_1) { +_start: { -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; size_t x_71; size_t x_72; size_t x_73; size_t x_74; size_t x_75; lean_object* x_76; uint8_t x_77; -x_61 = lean_ctor_get(x_10, 0); -x_62 = lean_ctor_get(x_10, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_10); -x_63 = lean_array_get_size(x_62); -x_64 = lean_uint64_of_nat(x_1); -x_65 = 32; -x_66 = lean_uint64_shift_right(x_64, x_65); -x_67 = lean_uint64_xor(x_64, x_66); -x_68 = 16; -x_69 = lean_uint64_shift_right(x_67, x_68); -x_70 = lean_uint64_xor(x_67, x_69); -x_71 = lean_uint64_to_usize(x_70); -x_72 = lean_usize_of_nat(x_63); -lean_dec(x_63); -x_73 = 1; -x_74 = lean_usize_sub(x_72, x_73); -x_75 = lean_usize_land(x_71, x_74); -x_76 = lean_array_uget(x_62, x_75); -x_77 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_76); -if (x_77 == 0) +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isPatternDontCare(lean_object* x_1) { +_start: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_78 = lean_unsigned_to_nat(1u); -x_79 = lean_nat_add(x_61, x_78); -lean_dec(x_61); -x_80 = lean_box(0); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_1); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_76); -x_82 = lean_array_uset(x_62, x_75, x_81); -x_83 = lean_unsigned_to_nat(4u); -x_84 = lean_nat_mul(x_79, x_83); -x_85 = lean_unsigned_to_nat(3u); -x_86 = lean_nat_div(x_84, x_85); -lean_dec(x_84); -x_87 = lean_array_get_size(x_82); -x_88 = lean_nat_dec_le(x_86, x_87); -lean_dec(x_87); -lean_dec(x_86); -if (x_88 == 0) +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_3 = lean_expr_eqv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isPatternDontCare___boxed(lean_object* x_1) { +_start: { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_89 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_82); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_79); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set(x_9, 2, x_90); -x_91 = lean_st_ref_set(x_2, x_9, x_11); -x_92 = lean_ctor_get(x_91, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_93 = x_91; -} else { - lean_dec_ref(x_91); - x_93 = lean_box(0); +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_isPatternDontCare(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; } -if (lean_is_scalar(x_93)) { - x_94 = lean_alloc_ctor(0, 2, 0); -} else { - x_94 = x_93; } -lean_ctor_set(x_94, 0, x_80); -lean_ctor_set(x_94, 1, x_92); -return x_94; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_Lean_Expr_isBVar(x_1); +if (x_2 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_4 = lean_expr_eqv(x_1, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isGroundPattern(x_1); +return x_5; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_79); -lean_ctor_set(x_95, 1, x_82); -lean_ctor_set(x_9, 2, x_95); -x_96 = lean_st_ref_set(x_2, x_9, x_11); -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_98 = x_96; -} else { - lean_dec_ref(x_96); - x_98 = lean_box(0); +uint8_t x_6; +x_6 = 1; +return x_6; } -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_98; } -lean_ctor_set(x_99, 0, x_80); -lean_ctor_set(x_99, 1, x_97); -return x_99; +else +{ +uint8_t x_7; +x_7 = 1; +return x_7; } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern___boxed(lean_object* x_1) { +_start: { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -lean_dec(x_76); +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_1); lean_dec(x_1); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_61); -lean_ctor_set(x_100, 1, x_62); -lean_ctor_set(x_9, 2, x_100); -x_101 = lean_st_ref_set(x_2, x_9, x_11); -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_103 = x_101; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; +x_3 = lean_box(x_2); +return x_3; } -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; } +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" ", 1, 1); +return x_1; } } -else +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2() { +_start: { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint64_t x_112; uint64_t x_113; uint64_t x_114; uint64_t x_115; uint64_t x_116; uint64_t x_117; uint64_t x_118; size_t x_119; size_t x_120; size_t x_121; size_t x_122; size_t x_123; lean_object* x_124; uint8_t x_125; -x_106 = lean_ctor_get(x_9, 0); -x_107 = lean_ctor_get(x_9, 1); -lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_9); -x_108 = lean_ctor_get(x_10, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_10, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - x_110 = x_10; -} else { - lean_dec_ref(x_10); - x_110 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -x_111 = lean_array_get_size(x_109); -x_112 = lean_uint64_of_nat(x_1); -x_113 = 32; -x_114 = lean_uint64_shift_right(x_112, x_113); -x_115 = lean_uint64_xor(x_112, x_114); -x_116 = 16; -x_117 = lean_uint64_shift_right(x_115, x_116); -x_118 = lean_uint64_xor(x_115, x_117); -x_119 = lean_uint64_to_usize(x_118); -x_120 = lean_usize_of_nat(x_111); -lean_dec(x_111); -x_121 = 1; -x_122 = lean_usize_sub(x_120, x_121); -x_123 = lean_usize_land(x_119, x_122); -x_124 = lean_array_uget(x_109, x_123); -x_125 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_124); -if (x_125 == 0) +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3() { +_start: { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_126 = lean_unsigned_to_nat(1u); -x_127 = lean_nat_add(x_108, x_126); -lean_dec(x_108); -x_128 = lean_box(0); -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_1); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_129, 2, x_124); -x_130 = lean_array_uset(x_109, x_123, x_129); -x_131 = lean_unsigned_to_nat(4u); -x_132 = lean_nat_mul(x_127, x_131); -x_133 = lean_unsigned_to_nat(3u); -x_134 = lean_nat_div(x_132, x_133); -lean_dec(x_132); -x_135 = lean_array_get_size(x_130); -x_136 = lean_nat_dec_le(x_134, x_135); -lean_dec(x_135); -lean_dec(x_134); -if (x_136 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_137 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_130); -if (lean_is_scalar(x_110)) { - x_138 = lean_alloc_ctor(0, 2, 0); -} else { - x_138 = x_110; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___closed__3; +x_5 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_1); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; } -lean_ctor_set(x_138, 0, x_127); -lean_ctor_set(x_138, 1, x_137); -x_139 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_139, 0, x_106); -lean_ctor_set(x_139, 1, x_107); -lean_ctor_set(x_139, 2, x_138); -x_140 = lean_st_ref_set(x_2, x_139, x_11); -x_141 = lean_ctor_get(x_140, 1); -lean_inc(x_141); -if (lean_is_exclusive(x_140)) { - lean_ctor_release(x_140, 0); - lean_ctor_release(x_140, 1); - x_142 = x_140; -} else { - lean_dec_ref(x_140); - x_142 = lean_box(0); } -if (lean_is_scalar(x_142)) { - x_143 = lean_alloc_ctor(0, 2, 0); -} else { - x_143 = x_142; +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed), 3, 0); +return x_1; } -lean_ctor_set(x_143, 0, x_128); -lean_ctor_set(x_143, 1, x_141); -return x_143; } -else +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2() { +_start: { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -if (lean_is_scalar(x_110)) { - x_144 = lean_alloc_ctor(0, 2, 0); -} else { - x_144 = x_110; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; } -lean_ctor_set(x_144, 0, x_127); -lean_ctor_set(x_144, 1, x_130); -x_145 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_145, 0, x_106); -lean_ctor_set(x_145, 1, x_107); -lean_ctor_set(x_145, 2, x_144); -x_146 = lean_st_ref_set(x_2, x_145, x_11); -x_147 = lean_ctor_get(x_146, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_146)) { - lean_ctor_release(x_146, 0); - lean_ctor_release(x_146, 1); - x_148 = x_146; -} else { - lean_dec_ref(x_146); - x_148 = lean_box(0); } -if (lean_is_scalar(x_148)) { - x_149 = lean_alloc_ctor(0, 2, 0); -} else { - x_149 = x_148; +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; } -lean_ctor_set(x_149, 0, x_128); -lean_ctor_set(x_149, 1, x_147); -return x_149; } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -lean_dec(x_124); -lean_dec(x_1); -if (lean_is_scalar(x_110)) { - x_150 = lean_alloc_ctor(0, 2, 0); -} else { - x_150 = x_110; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } -lean_ctor_set(x_150, 0, x_108); -lean_ctor_set(x_150, 1, x_109); -x_151 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_151, 0, x_106); -lean_ctor_set(x_151, 1, x_107); -lean_ctor_set(x_151, 2, x_150); -x_152 = lean_st_ref_set(x_2, x_151, x_11); -x_153 = lean_ctor_get(x_152, 1); -lean_inc(x_153); -if (lean_is_exclusive(x_152)) { - lean_ctor_release(x_152, 0); - lean_ctor_release(x_152, 1); - x_154 = x_152; -} else { - lean_dec_ref(x_152); - x_154 = lean_box(0); +else +{ +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } -x_155 = lean_box(0); -if (lean_is_scalar(x_154)) { - x_156 = lean_alloc_ctor(0, 2, 0); -} else { - x_156 = x_154; } -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_153); -return x_156; +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -uint8_t x_2; -x_2 = l_Lean_Expr_isApp(x_1); -if (x_2 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; +return x_6; } else { -lean_object* x_4; -x_4 = l_Lean_Expr_getAppFn(x_1); -switch (lean_obj_tag(x_4)) { -case 1: -{ -lean_object* x_5; -x_5 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_5, 0, x_4); -return x_5; -} -case 4: +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); -x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; -x_8 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_7, x_6); -lean_dec(x_6); -if (x_8 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_9; -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_4); -return x_9; +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -lean_object* x_10; -lean_dec(x_4); -x_10 = lean_box(0); -return x_10; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } -default: +else { -lean_object* x_11; -lean_dec(x_4); -x_11 = lean_box(0); -return x_11; -} +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); -lean_dec(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_2, x_1); -if (x_9 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_10; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_3); -lean_ctor_set(x_10, 1, x_8); -return x_10; +return x_6; } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; uint8_t x_22; lean_object* x_23; lean_object* x_38; -x_11 = lean_array_uget(x_3, x_2); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_array_uset(x_3, x_2, x_12); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_11); -x_38 = l_Lean_Meta_isTypeFormer(x_11, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_38) == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_39; uint8_t x_40; -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_unbox(x_39); -if (x_40 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_41; lean_object* x_42; -lean_dec(x_39); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -lean_dec(x_38); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_11); -x_42 = l_Lean_Meta_isProof(x_11, x_4, x_5, x_6, x_7, x_41); -if (lean_obj_tag(x_42) == 0) +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; +} +else { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = lean_unbox(x_43); -lean_dec(x_43); -x_22 = x_45; -x_23 = x_44; -goto block_37; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} } else { -uint8_t x_46; -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_46 = !lean_is_exclusive(x_42); -if (x_46 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -return x_42; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_42, 0); -x_48 = lean_ctor_get(x_42, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_42); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -else -{ -lean_object* x_50; uint8_t x_51; -x_50 = lean_ctor_get(x_38, 1); -lean_inc(x_50); -lean_dec(x_38); -x_51 = lean_unbox(x_39); -lean_dec(x_39); -x_22 = x_51; -x_23 = x_50; -goto block_37; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: { -uint8_t x_52; -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_52 = !lean_is_exclusive(x_38); -if (x_52 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -return x_38; +return x_6; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_38, 0); -x_54 = lean_ctor_get(x_38, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_38); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; -} -} -block_21: -{ -size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; -x_16 = 1; -x_17 = lean_usize_add(x_2, x_16); -x_18 = lean_box(x_14); -x_19 = lean_array_uset(x_13, x_2, x_18); -x_2 = x_17; -x_3 = x_19; -x_8 = x_15; -goto _start; -} -block_37: -{ -if (x_22 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = l_Lean_Expr_fvarId_x21(x_11); -lean_dec(x_11); -lean_inc(x_4); -x_25 = l_Lean_FVarId_getDecl(x_24, x_4, x_5, x_6, x_7, x_23); -if (lean_obj_tag(x_25) == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = l_Lean_LocalDecl_binderInfo(x_26); -lean_dec(x_26); -x_29 = lean_box(x_28); -if (lean_obj_tag(x_29) == 3) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_30; -x_30 = 1; -x_14 = x_30; -x_15 = x_27; -goto block_21; +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -uint8_t x_31; -lean_dec(x_29); -x_31 = 0; -x_14 = x_31; -x_15 = x_27; -goto block_21; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } else { -uint8_t x_32; -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_32 = !lean_is_exclusive(x_25); -if (x_32 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -return x_25; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_25, 0); -x_34 = lean_ctor_get(x_25, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_25); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } -else -{ -uint8_t x_36; -lean_dec(x_11); -x_36 = 1; -x_14 = x_36; -x_15 = x_23; -goto block_21; } } } -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_array_size(x_1); -x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(x_8, x_9, x_1, x_3, x_4, x_5, x_6, x_7); -return x_10; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1() { -_start: +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed), 7, 0); -return x_1; -} +return x_6; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); lean_dec(x_8); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_2); -x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1; -x_13 = 0; -x_14 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_9, x_11, x_12, x_13, x_3, x_4, x_5, x_6, x_10); -return x_14; +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -uint8_t x_15; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_15 = !lean_is_exclusive(x_8); -if (x_15 == 0) +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} +} +else { -return x_8; +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_8, 0); -x_17 = lean_ctor_get(x_8, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_8); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_10 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___spec__1(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; } +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_5, 5); -x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_11); -lean_ctor_set_tag(x_9, 1); -lean_ctor_set(x_9, 0, x_12); -return x_9; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_9); -lean_inc(x_8); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_array_set(x_2, x_1, x_3); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_9); -return x_12; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_3, 1); -x_15 = lean_nat_dec_lt(x_5, x_14); -if (x_15 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_16; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_5); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_4); -lean_ctor_set(x_16, 1, x_13); -return x_16; +return x_6; } else { -lean_object* x_17; lean_object* x_35; uint8_t x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; -x_35 = lean_array_get_size(x_4); -x_36 = lean_nat_dec_lt(x_5, x_35); -lean_dec(x_35); -x_37 = lean_array_get_size(x_1); -x_38 = lean_nat_dec_lt(x_5, x_37); -lean_dec(x_37); -if (x_36 == 0) -{ -lean_object* x_93; lean_object* x_94; -x_93 = l_Lean_instInhabitedExpr; -x_94 = l_outOfBounds___rarg(x_93); -x_39 = x_94; -goto block_92; -} -else +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_95; -x_95 = lean_array_fget(x_4, x_5); -x_39 = x_95; -goto block_92; -} -block_34: +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_5); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_18, 0); -lean_inc(x_21); -lean_dec(x_18); -lean_ctor_set(x_17, 0, x_21); +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); return x_17; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_ctor_get(x_18, 0); -lean_inc(x_23); -lean_dec(x_18); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_17, 1); -lean_inc(x_25); -lean_dec(x_17); -x_26 = lean_ctor_get(x_18, 0); -lean_inc(x_26); -lean_dec(x_18); -x_27 = lean_ctor_get(x_3, 2); -x_28 = lean_nat_add(x_5, x_27); -lean_dec(x_5); -x_4 = x_26; -x_5 = x_28; -x_6 = lean_box(0); -x_7 = lean_box(0); -x_13 = x_25; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; goto _start; } } else { -uint8_t x_30; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_5); -x_30 = !lean_is_exclusive(x_17); -if (x_30 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -return x_17; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_17, 0); -x_32 = lean_ctor_get(x_17, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_17); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; +} +} } } } -block_92: +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: { -uint8_t x_40; uint8_t x_41; -x_40 = l_Lean_Expr_hasLooseBVars(x_39); -if (x_38 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -uint8_t x_89; -x_89 = 0; -x_41 = x_89; -goto block_88; +return x_6; } else { -lean_object* x_90; uint8_t x_91; -x_90 = lean_array_fget(x_1, x_5); -x_91 = lean_unbox(x_90); -lean_dec(x_90); -x_41 = x_91; -goto block_88; -} -block_88: -{ -if (x_40 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -uint8_t x_42; -x_42 = l_Lean_Expr_hasMVar(x_39); -if (x_42 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_43 = l_Lean_Meta_Grind_mkGroundPattern(x_39); -x_44 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_43, x_8, x_9, x_10, x_11, x_12, x_13); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = lean_ctor_get(x_45, 0); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_ctor_get(x_3, 2); -x_49 = lean_nat_add(x_5, x_48); -lean_dec(x_5); -x_4 = x_47; -x_5 = x_49; -x_6 = lean_box(0); -x_7 = lean_box(0); -x_13 = x_46; -goto _start; +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_dec(x_39); -x_51 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_52 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_51, x_8, x_9, x_10, x_11, x_12, x_13); -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); -lean_dec(x_52); -x_55 = lean_ctor_get(x_53, 0); -lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_ctor_get(x_3, 2); -x_57 = lean_nat_add(x_5, x_56); -lean_dec(x_5); -x_4 = x_55; -x_5 = x_57; -x_6 = lean_box(0); -x_7 = lean_box(0); -x_13 = x_54; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; goto _start; } } else { -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_39, 0); -lean_inc(x_59); -x_60 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_13); -if (x_41 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_62 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_61); -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -lean_dec(x_62); -x_64 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_39, x_8, x_9, x_10, x_11, x_12, x_63); -x_17 = x_64; -goto block_34; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_65; uint8_t x_66; -x_65 = lean_ctor_get(x_60, 0); -lean_inc(x_65); -x_66 = lean_unbox(x_65); -lean_dec(x_65); -if (x_66 == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_67 = lean_ctor_get(x_60, 1); -lean_inc(x_67); -lean_dec(x_60); -x_68 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_59, x_8, x_9, x_10, x_11, x_12, x_67); -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -lean_dec(x_68); -x_70 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_39, x_8, x_9, x_10, x_11, x_12, x_69); -x_17 = x_70; -goto block_34; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; -lean_dec(x_59); -lean_dec(x_39); -x_71 = lean_ctor_get(x_60, 1); -lean_inc(x_71); -lean_dec(x_60); -x_72 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_73 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_72, x_8, x_9, x_10, x_11, x_12, x_71); -x_17 = x_73; -goto block_34; } } } -else -{ -if (x_41 == 0) +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: { -lean_object* x_74; -x_74 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_39); -if (lean_obj_tag(x_74) == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_75; lean_object* x_76; -lean_dec(x_39); -x_75 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_75, x_8, x_9, x_10, x_11, x_12, x_13); -x_17 = x_76; -goto block_34; +return x_6; } else { -uint8_t x_77; lean_object* x_78; -lean_dec(x_74); -x_77 = 0; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); lean_inc(x_8); -x_78 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_39, x_77, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_78) == 0) +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec(x_78); -x_81 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_79, x_8, x_9, x_10, x_11, x_12, x_80); -x_17 = x_81; -goto block_34; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } else { -uint8_t x_82; -lean_dec(x_4); -x_82 = !lean_is_exclusive(x_78); -if (x_82 == 0) -{ -x_17 = x_78; -goto block_34; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; +} } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_78, 0); -x_84 = lean_ctor_get(x_78, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_78); -x_85 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -x_17 = x_85; -goto block_34; +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: { -lean_object* x_86; lean_object* x_87; -lean_dec(x_39); -x_86 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_87 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_5, x_4, x_86, x_8, x_9, x_10, x_11, x_12, x_13); -x_17 = x_87; -goto block_34; +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; } +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } +else +{ +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } +else +{ +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; } } -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2() { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; -x_2 = l_Lean_instInhabitedExpr; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; } +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +lean_inc(x_8); +x_9 = l_Lean_Meta_Grind_ppPattern(x_8); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__1; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isAtomicPattern(x_8); +lean_dec(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__2; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___closed__3; +x_14 = l_Lean_MessageData_bracket(x_12, x_9, x_13); +x_15 = lean_box(0); +x_16 = lean_apply_3(x_10, x_14, x_6, x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +return x_17; } -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; -x_9 = lean_panic_fn(x_8, x_1); -x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -return x_10; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_5 = x_20; +x_6 = x_18; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_5, 5); -x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_apply_3(x_10, x_9, x_6, x_22); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_11); -lean_ctor_set_tag(x_9, 1); -lean_ctor_set(x_9, 0, x_12); -return x_9; +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_9); -lean_inc(x_8); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; +lean_object* x_25; size_t x_26; size_t x_27; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = 1; +x_27 = lean_usize_add(x_5, x_26); +x_5 = x_27; +x_6 = x_25; +goto _start; } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1() { +} +} +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid pattern, (non-forbidden) application expected", 53, 53); +x_1 = lean_mk_string_unchecked("#", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("f.isConst || f.isFVar\n ", 24, 24); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.NormalizePattern.go", 83, 83); +x_1 = lean_mk_string_unchecked("\?", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; -x_3 = lean_unsigned_to_nat(177u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__8() { _start: { -lean_object* x_9; -x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; -lean_dec(x_1); -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; -x_11 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_10, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_11; -} -else -{ -lean_object* x_12; lean_object* x_13; uint8_t x_46; -x_12 = lean_ctor_get(x_9, 0); -lean_inc(x_12); -lean_dec(x_9); -x_46 = l_Lean_Expr_isConst(x_12); -if (x_46 == 0) -{ -uint8_t x_47; -x_47 = l_Lean_Expr_isFVar(x_12); -if (x_47 == 0) -{ -lean_object* x_48; lean_object* x_49; -lean_dec(x_12); -lean_dec(x_1); -x_48 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; -x_49 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(x_48, x_3, x_4, x_5, x_6, x_7, x_8); -return x_49; -} -else -{ -lean_object* x_50; -x_50 = lean_box(0); -x_13 = x_50; -goto block_45; -} -} -else -{ -lean_object* x_51; -x_51 = lean_box(0); -x_13 = x_51; -goto block_45; -} -block_45: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_dec(x_13); -lean_inc(x_12); -x_14 = l_Lean_Expr_toHeadIndex(x_12); -x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_14, x_3, x_4, x_5, x_6, x_7, x_8); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_unsigned_to_nat(0u); -x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_17); -x_19 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_18); -x_20 = lean_mk_array(x_18, x_19); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_18, x_21); -lean_dec(x_18); -x_23 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_20, x_22); -x_24 = lean_array_get_size(x_23); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_24); -lean_inc(x_12); -x_25 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask(x_12, x_24, x_4, x_5, x_6, x_7, x_16); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_28, 0, x_17); -lean_ctor_set(x_28, 1, x_24); -lean_ctor_set(x_28, 2, x_21); -x_29 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_26, x_28, x_28, x_23, x_17, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_7, x_27); -lean_dec(x_28); -lean_dec(x_26); -if (lean_obj_tag(x_29) == 0) -{ -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_29, 0); -x_32 = l_Lean_mkAppN(x_12, x_31); -lean_dec(x_31); -lean_ctor_set(x_29, 0, x_32); -return x_29; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_ctor_get(x_29, 0); -x_34 = lean_ctor_get(x_29, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_29); -x_35 = l_Lean_mkAppN(x_12, x_33); -lean_dec(x_33); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -return x_36; -} -} -else -{ -uint8_t x_37; -lean_dec(x_12); -x_37 = !lean_is_exclusive(x_29); -if (x_37 == 0) -{ -return x_29; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_29, 0); -x_39 = lean_ctor_get(x_29, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_29); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`[", 2, 2); +return x_1; } -else -{ -uint8_t x_41; -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_41 = !lean_is_exclusive(x_25); -if (x_41 == 0) -{ -return x_25; } -else +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__9() { +_start: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_25, 0); -x_43 = lean_ctor_get(x_25, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_25); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; -} -} -} -} +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__10() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid pattern, it does not have pattern variables", 51, 51); +x_1 = lean_mk_string_unchecked("]", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_ppPattern___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; +x_1 = l_Lean_Meta_Grind_ppPattern___closed__10; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ppPattern(lean_object* x_1) { _start: { -if (x_2 == 0) +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +if (lean_obj_tag(x_2) == 0) { -lean_object* x_9; lean_object* x_10; -x_9 = lean_box(0); -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -else +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_4 = lean_expr_eqv(x_1, x_3); +if (x_4 == 0) { -uint8_t x_11; -x_11 = l_Lean_Expr_hasLooseBVars(x_1); -if (x_11 == 0) +switch (lean_obj_tag(x_1)) { +case 0: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); lean_dec(x_1); -x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; -x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_12, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -return x_13; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_13); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -} -else -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_box(0); -x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8); -return x_19; -} -} -} +x_6 = l___private_Init_Data_Repr_0__Nat_reprFast(x_5); +x_7 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = l_Lean_MessageData_ofFormat(x_7); +x_9 = l_Lean_Meta_Grind_ppPattern___closed__2; +x_10 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +case 1: { -lean_object* x_8; -x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; +x_13 = l_Lean_Expr_getAppFn(x_1); +x_14 = l_Lean_MessageData_ofExpr(x_13); +x_15 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = lean_unsigned_to_nat(0u); +x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_18); +x_20 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_19); +x_21 = lean_mk_array(x_19, x_20); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_sub(x_19, x_22); +lean_dec(x_19); +x_24 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_21, x_23); +x_25 = lean_box(0); +x_26 = lean_array_size(x_24); +x_27 = 0; +x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_24, x_25, x_24, x_26, x_27, x_17); +lean_dec(x_24); +return x_28; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +case 2: { -lean_object* x_10; -x_10 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; size_t x_43; lean_object* x_44; +x_29 = l_Lean_Expr_getAppFn(x_1); +x_30 = l_Lean_MessageData_ofExpr(x_29); +x_31 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = lean_unsigned_to_nat(0u); +x_35 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_34); +x_36 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_35); +x_37 = lean_mk_array(x_35, x_36); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_nat_sub(x_35, x_38); +lean_dec(x_35); +x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_37, x_39); +x_41 = lean_box(0); +x_42 = lean_array_size(x_40); +x_43 = 0; +x_44 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_40, x_41, x_40, x_42, x_43, x_33); +lean_dec(x_40); +return x_44; +} +case 3: +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; +x_45 = l_Lean_Expr_getAppFn(x_1); +x_46 = l_Lean_MessageData_ofExpr(x_45); +x_47 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = lean_unsigned_to_nat(0u); +x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); +x_52 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_51); +x_53 = lean_mk_array(x_51, x_52); +x_54 = lean_unsigned_to_nat(1u); +x_55 = lean_nat_sub(x_51, x_54); +lean_dec(x_51); +x_56 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_53, x_55); +x_57 = lean_box(0); +x_58 = lean_array_size(x_56); +x_59 = 0; +x_60 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_56, x_57, x_56, x_58, x_59, x_49); +lean_dec(x_56); +return x_60; +} +case 4: +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; +x_61 = l_Lean_Expr_getAppFn(x_1); +x_62 = l_Lean_MessageData_ofExpr(x_61); +x_63 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_64 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_62); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +x_66 = lean_unsigned_to_nat(0u); +x_67 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_66); +x_68 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_67); +x_69 = lean_mk_array(x_67, x_68); +x_70 = lean_unsigned_to_nat(1u); +x_71 = lean_nat_sub(x_67, x_70); +lean_dec(x_67); +x_72 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_69, x_71); +x_73 = lean_box(0); +x_74 = lean_array_size(x_72); +x_75 = 0; +x_76 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_72, x_73, x_72, x_74, x_75, x_65); +lean_dec(x_72); +return x_76; +} +case 5: +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; size_t x_90; size_t x_91; lean_object* x_92; +x_77 = l_Lean_Expr_getAppFn(x_1); +x_78 = l_Lean_MessageData_ofExpr(x_77); +x_79 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_78); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +x_82 = lean_unsigned_to_nat(0u); +x_83 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_82); +x_84 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_83); +x_85 = lean_mk_array(x_83, x_84); +x_86 = lean_unsigned_to_nat(1u); +x_87 = lean_nat_sub(x_83, x_86); +lean_dec(x_83); +x_88 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_85, x_87); +x_89 = lean_box(0); +x_90 = lean_array_size(x_88); +x_91 = 0; +x_92 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_88, x_89, x_88, x_90, x_91, x_81); +lean_dec(x_88); +return x_92; +} +case 6: +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; size_t x_106; size_t x_107; lean_object* x_108; +x_93 = l_Lean_Expr_getAppFn(x_1); +x_94 = l_Lean_MessageData_ofExpr(x_93); +x_95 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_96 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_94); +x_97 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = lean_unsigned_to_nat(0u); +x_99 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_98); +x_100 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_99); +x_101 = lean_mk_array(x_99, x_100); +x_102 = lean_unsigned_to_nat(1u); +x_103 = lean_nat_sub(x_99, x_102); +lean_dec(x_99); +x_104 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_101, x_103); +x_105 = lean_box(0); +x_106 = lean_array_size(x_104); +x_107 = 0; +x_108 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_104, x_105, x_104, x_106, x_107, x_97); +lean_dec(x_104); +return x_108; +} +case 7: +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; size_t x_122; size_t x_123; lean_object* x_124; +x_109 = l_Lean_Expr_getAppFn(x_1); +x_110 = l_Lean_MessageData_ofExpr(x_109); +x_111 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_112 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_110); +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_111); +x_114 = lean_unsigned_to_nat(0u); +x_115 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_114); +x_116 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_115); +x_117 = lean_mk_array(x_115, x_116); +x_118 = lean_unsigned_to_nat(1u); +x_119 = lean_nat_sub(x_115, x_118); +lean_dec(x_115); +x_120 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_117, x_119); +x_121 = lean_box(0); +x_122 = lean_array_size(x_120); +x_123 = 0; +x_124 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_120, x_121, x_120, x_122, x_123, x_113); +lean_dec(x_120); +return x_124; +} +case 8: +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; size_t x_138; size_t x_139; lean_object* x_140; +x_125 = l_Lean_Expr_getAppFn(x_1); +x_126 = l_Lean_MessageData_ofExpr(x_125); +x_127 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_128 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_126); +x_129 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_129, 0, x_128); +lean_ctor_set(x_129, 1, x_127); +x_130 = lean_unsigned_to_nat(0u); +x_131 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_130); +x_132 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_131); +x_133 = lean_mk_array(x_131, x_132); +x_134 = lean_unsigned_to_nat(1u); +x_135 = lean_nat_sub(x_131, x_134); +lean_dec(x_131); +x_136 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_133, x_135); +x_137 = lean_box(0); +x_138 = lean_array_size(x_136); +x_139 = 0; +x_140 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_136, x_137, x_136, x_138, x_139, x_129); +lean_dec(x_136); +return x_140; +} +case 9: +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; size_t x_154; size_t x_155; lean_object* x_156; +x_141 = l_Lean_Expr_getAppFn(x_1); +x_142 = l_Lean_MessageData_ofExpr(x_141); +x_143 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_144 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_142); +x_145 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_143); +x_146 = lean_unsigned_to_nat(0u); +x_147 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_146); +x_148 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_147); +x_149 = lean_mk_array(x_147, x_148); +x_150 = lean_unsigned_to_nat(1u); +x_151 = lean_nat_sub(x_147, x_150); +lean_dec(x_147); +x_152 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_149, x_151); +x_153 = lean_box(0); +x_154 = lean_array_size(x_152); +x_155 = 0; +x_156 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_152, x_153, x_152, x_154, x_155, x_145); +lean_dec(x_152); +return x_156; +} +case 10: +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; size_t x_170; size_t x_171; lean_object* x_172; +x_157 = l_Lean_Expr_getAppFn(x_1); +x_158 = l_Lean_MessageData_ofExpr(x_157); +x_159 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_160 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_158); +x_161 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_159); +x_162 = lean_unsigned_to_nat(0u); +x_163 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_162); +x_164 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_163); +x_165 = lean_mk_array(x_163, x_164); +x_166 = lean_unsigned_to_nat(1u); +x_167 = lean_nat_sub(x_163, x_166); +lean_dec(x_163); +x_168 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_165, x_167); +x_169 = lean_box(0); +x_170 = lean_array_size(x_168); +x_171 = 0; +x_172 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_168, x_169, x_168, x_170, x_171, x_161); +lean_dec(x_168); +return x_172; +} +default: +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; size_t x_186; size_t x_187; lean_object* x_188; +x_173 = l_Lean_Expr_getAppFn(x_1); +x_174 = l_Lean_MessageData_ofExpr(x_173); +x_175 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_176 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_176, 0, x_175); +lean_ctor_set(x_176, 1, x_174); +x_177 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_175); +x_178 = lean_unsigned_to_nat(0u); +x_179 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_178); +x_180 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_179); +x_181 = lean_mk_array(x_179, x_180); +x_182 = lean_unsigned_to_nat(1u); +x_183 = lean_nat_sub(x_179, x_182); +lean_dec(x_179); +x_184 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_181, x_183); +x_185 = lean_box(0); +x_186 = lean_array_size(x_184); +x_187 = 0; +x_188 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_184, x_185, x_184, x_186, x_187, x_177); +lean_dec(x_184); +return x_188; +} +} +} +else +{ +lean_object* x_189; +lean_dec(x_1); +x_189 = l_Lean_Meta_Grind_ppPattern___closed__7; +return x_189; +} +} +else +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +lean_dec(x_1); +x_190 = lean_ctor_get(x_2, 0); +lean_inc(x_190); +lean_dec(x_2); +x_191 = l_Lean_MessageData_ofExpr(x_190); +x_192 = l_Lean_Meta_Grind_ppPattern___closed__9; +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_191); +x_194 = l_Lean_Meta_Grind_ppPattern___closed__11; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +return x_195; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__1(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__2(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__3(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__4(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__5(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__7(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__8(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__9(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_ppPattern___spec__11(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = l_Lean_HeadIndex_hash(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = l_Lean_HeadIndex_hash(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__4(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__3(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; lean_object* x_28; uint8_t x_29; +x_12 = lean_ctor_get(x_8, 1); +x_13 = lean_ctor_get(x_8, 0); +lean_dec(x_13); +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_array_get_size(x_14); +x_16 = l_Lean_HeadIndex_hash(x_1); +x_17 = 32; +x_18 = lean_uint64_shift_right(x_16, x_17); +x_19 = lean_uint64_xor(x_16, x_18); +x_20 = 16; +x_21 = lean_uint64_shift_right(x_19, x_20); +x_22 = lean_uint64_xor(x_19, x_21); +x_23 = lean_uint64_to_usize(x_22); +x_24 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_25 = 1; +x_26 = lean_usize_sub(x_24, x_25); +x_27 = lean_usize_land(x_23, x_26); +x_28 = lean_array_uget(x_14, x_27); +lean_dec(x_14); +x_29 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_28); +lean_dec(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_free_object(x_8); +x_30 = lean_st_ref_take(x_2, x_12); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = lean_ctor_get(x_31, 0); +x_35 = lean_ctor_get(x_31, 1); +lean_inc(x_1); +x_36 = lean_array_push(x_34, x_1); +x_37 = !lean_is_exclusive(x_35); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; uint8_t x_45; +x_38 = lean_ctor_get(x_35, 0); +x_39 = lean_ctor_get(x_35, 1); +x_40 = lean_array_get_size(x_39); +x_41 = lean_usize_of_nat(x_40); +lean_dec(x_40); +x_42 = lean_usize_sub(x_41, x_25); +x_43 = lean_usize_land(x_23, x_42); +x_44 = lean_array_uget(x_39, x_43); +x_45 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_44); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_38, x_46); +lean_dec(x_38); +x_48 = lean_box(0); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_1); +lean_ctor_set(x_49, 1, x_48); +lean_ctor_set(x_49, 2, x_44); +x_50 = lean_array_uset(x_39, x_43, x_49); +x_51 = lean_unsigned_to_nat(4u); +x_52 = lean_nat_mul(x_47, x_51); +x_53 = lean_unsigned_to_nat(3u); +x_54 = lean_nat_div(x_52, x_53); +lean_dec(x_52); +x_55 = lean_array_get_size(x_50); +x_56 = lean_nat_dec_le(x_54, x_55); +lean_dec(x_55); +lean_dec(x_54); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_57 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_50); +lean_ctor_set(x_35, 1, x_57); +lean_ctor_set(x_35, 0, x_47); +lean_ctor_set(x_31, 0, x_36); +x_58 = lean_st_ref_set(x_2, x_31, x_32); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +lean_object* x_60; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +lean_ctor_set(x_58, 0, x_48); +return x_58; +} +else +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_48); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +else +{ +lean_object* x_63; uint8_t x_64; +lean_ctor_set(x_35, 1, x_50); +lean_ctor_set(x_35, 0, x_47); +lean_ctor_set(x_31, 0, x_36); +x_63 = lean_st_ref_set(x_2, x_31, x_32); +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) +{ +lean_object* x_65; +x_65 = lean_ctor_get(x_63, 0); +lean_dec(x_65); +lean_ctor_set(x_63, 0, x_48); +return x_63; +} +else +{ +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_dec(x_63); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_48); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +lean_object* x_68; uint8_t x_69; +lean_dec(x_44); +lean_dec(x_1); +lean_ctor_set(x_31, 0, x_36); +x_68 = lean_st_ref_set(x_2, x_31, x_32); +x_69 = !lean_is_exclusive(x_68); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_68, 0); +lean_dec(x_70); +x_71 = lean_box(0); +lean_ctor_set(x_68, 0, x_71); +return x_68; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_68, 1); +lean_inc(x_72); +lean_dec(x_68); +x_73 = lean_box(0); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_72); +return x_74; +} +} +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; size_t x_80; lean_object* x_81; uint8_t x_82; +x_75 = lean_ctor_get(x_35, 0); +x_76 = lean_ctor_get(x_35, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_35); +x_77 = lean_array_get_size(x_76); +x_78 = lean_usize_of_nat(x_77); +lean_dec(x_77); +x_79 = lean_usize_sub(x_78, x_25); +x_80 = lean_usize_land(x_23, x_79); +x_81 = lean_array_uget(x_76, x_80); +x_82 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_81); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_83 = lean_unsigned_to_nat(1u); +x_84 = lean_nat_add(x_75, x_83); +lean_dec(x_75); +x_85 = lean_box(0); +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_1); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_81); +x_87 = lean_array_uset(x_76, x_80, x_86); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_mul(x_84, x_88); +x_90 = lean_unsigned_to_nat(3u); +x_91 = lean_nat_div(x_89, x_90); +lean_dec(x_89); +x_92 = lean_array_get_size(x_87); +x_93 = lean_nat_dec_le(x_91, x_92); +lean_dec(x_92); +lean_dec(x_91); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_94 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_87); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_84); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_31, 1, x_95); +lean_ctor_set(x_31, 0, x_36); +x_96 = lean_st_ref_set(x_2, x_31, x_32); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_98 = x_96; +} else { + lean_dec_ref(x_96); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_85); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_84); +lean_ctor_set(x_100, 1, x_87); +lean_ctor_set(x_31, 1, x_100); +lean_ctor_set(x_31, 0, x_36); +x_101 = lean_st_ref_set(x_2, x_31, x_32); +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_103 = x_101; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); +} +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(0, 2, 0); +} else { + x_104 = x_103; +} +lean_ctor_set(x_104, 0, x_85); +lean_ctor_set(x_104, 1, x_102); +return x_104; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_81); +lean_dec(x_1); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_75); +lean_ctor_set(x_105, 1, x_76); +lean_ctor_set(x_31, 1, x_105); +lean_ctor_set(x_31, 0, x_36); +x_106 = lean_st_ref_set(x_2, x_31, x_32); +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_108 = x_106; +} else { + lean_dec_ref(x_106); + x_108 = lean_box(0); +} +x_109 = lean_box(0); +if (lean_is_scalar(x_108)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_108; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_107); +return x_110; +} +} +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; size_t x_119; size_t x_120; size_t x_121; lean_object* x_122; uint8_t x_123; +x_111 = lean_ctor_get(x_31, 0); +x_112 = lean_ctor_get(x_31, 1); +x_113 = lean_ctor_get(x_31, 2); +lean_inc(x_113); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_31); +lean_inc(x_1); +x_114 = lean_array_push(x_111, x_1); +x_115 = lean_ctor_get(x_112, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_112, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_117 = x_112; +} else { + lean_dec_ref(x_112); + x_117 = lean_box(0); +} +x_118 = lean_array_get_size(x_116); +x_119 = lean_usize_of_nat(x_118); +lean_dec(x_118); +x_120 = lean_usize_sub(x_119, x_25); +x_121 = lean_usize_land(x_23, x_120); +x_122 = lean_array_uget(x_116, x_121); +x_123 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_122); +if (x_123 == 0) +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; +x_124 = lean_unsigned_to_nat(1u); +x_125 = lean_nat_add(x_115, x_124); +lean_dec(x_115); +x_126 = lean_box(0); +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_1); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_122); +x_128 = lean_array_uset(x_116, x_121, x_127); +x_129 = lean_unsigned_to_nat(4u); +x_130 = lean_nat_mul(x_125, x_129); +x_131 = lean_unsigned_to_nat(3u); +x_132 = lean_nat_div(x_130, x_131); +lean_dec(x_130); +x_133 = lean_array_get_size(x_128); +x_134 = lean_nat_dec_le(x_132, x_133); +lean_dec(x_133); +lean_dec(x_132); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_135 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_128); +if (lean_is_scalar(x_117)) { + x_136 = lean_alloc_ctor(0, 2, 0); +} else { + x_136 = x_117; +} +lean_ctor_set(x_136, 0, x_125); +lean_ctor_set(x_136, 1, x_135); +x_137 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_137, 0, x_114); +lean_ctor_set(x_137, 1, x_136); +lean_ctor_set(x_137, 2, x_113); +x_138 = lean_st_ref_set(x_2, x_137, x_32); +x_139 = lean_ctor_get(x_138, 1); +lean_inc(x_139); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + x_140 = x_138; +} else { + lean_dec_ref(x_138); + x_140 = lean_box(0); +} +if (lean_is_scalar(x_140)) { + x_141 = lean_alloc_ctor(0, 2, 0); +} else { + x_141 = x_140; +} +lean_ctor_set(x_141, 0, x_126); +lean_ctor_set(x_141, 1, x_139); +return x_141; +} +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +if (lean_is_scalar(x_117)) { + x_142 = lean_alloc_ctor(0, 2, 0); +} else { + x_142 = x_117; +} +lean_ctor_set(x_142, 0, x_125); +lean_ctor_set(x_142, 1, x_128); +x_143 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_143, 0, x_114); +lean_ctor_set(x_143, 1, x_142); +lean_ctor_set(x_143, 2, x_113); +x_144 = lean_st_ref_set(x_2, x_143, x_32); +x_145 = lean_ctor_get(x_144, 1); +lean_inc(x_145); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_146 = x_144; +} else { + lean_dec_ref(x_144); + x_146 = lean_box(0); +} +if (lean_is_scalar(x_146)) { + x_147 = lean_alloc_ctor(0, 2, 0); +} else { + x_147 = x_146; +} +lean_ctor_set(x_147, 0, x_126); +lean_ctor_set(x_147, 1, x_145); +return x_147; +} +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +lean_dec(x_122); +lean_dec(x_1); +if (lean_is_scalar(x_117)) { + x_148 = lean_alloc_ctor(0, 2, 0); +} else { + x_148 = x_117; +} +lean_ctor_set(x_148, 0, x_115); +lean_ctor_set(x_148, 1, x_116); +x_149 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_149, 0, x_114); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set(x_149, 2, x_113); +x_150 = lean_st_ref_set(x_2, x_149, x_32); +x_151 = lean_ctor_get(x_150, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + lean_ctor_release(x_150, 1); + x_152 = x_150; +} else { + lean_dec_ref(x_150); + x_152 = lean_box(0); +} +x_153 = lean_box(0); +if (lean_is_scalar(x_152)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_152; +} +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_151); +return x_154; +} +} +} +else +{ +lean_object* x_155; +lean_dec(x_1); +x_155 = lean_box(0); +lean_ctor_set(x_8, 0, x_155); +return x_8; +} +} +else +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; uint64_t x_159; uint64_t x_160; uint64_t x_161; uint64_t x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; size_t x_166; size_t x_167; size_t x_168; size_t x_169; size_t x_170; lean_object* x_171; uint8_t x_172; +x_156 = lean_ctor_get(x_8, 1); +lean_inc(x_156); +lean_dec(x_8); +x_157 = lean_ctor_get(x_10, 1); +lean_inc(x_157); +lean_dec(x_10); +x_158 = lean_array_get_size(x_157); +x_159 = l_Lean_HeadIndex_hash(x_1); +x_160 = 32; +x_161 = lean_uint64_shift_right(x_159, x_160); +x_162 = lean_uint64_xor(x_159, x_161); +x_163 = 16; +x_164 = lean_uint64_shift_right(x_162, x_163); +x_165 = lean_uint64_xor(x_162, x_164); +x_166 = lean_uint64_to_usize(x_165); +x_167 = lean_usize_of_nat(x_158); +lean_dec(x_158); +x_168 = 1; +x_169 = lean_usize_sub(x_167, x_168); +x_170 = lean_usize_land(x_166, x_169); +x_171 = lean_array_uget(x_157, x_170); +lean_dec(x_157); +x_172 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_171); +lean_dec(x_171); +if (x_172 == 0) +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; size_t x_185; size_t x_186; size_t x_187; lean_object* x_188; uint8_t x_189; +x_173 = lean_st_ref_take(x_2, x_156); +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +lean_dec(x_173); +x_176 = lean_ctor_get(x_174, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_174, 1); +lean_inc(x_177); +x_178 = lean_ctor_get(x_174, 2); +lean_inc(x_178); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + lean_ctor_release(x_174, 2); + x_179 = x_174; +} else { + lean_dec_ref(x_174); + x_179 = lean_box(0); +} +lean_inc(x_1); +x_180 = lean_array_push(x_176, x_1); +x_181 = lean_ctor_get(x_177, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_177, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_183 = x_177; +} else { + lean_dec_ref(x_177); + x_183 = lean_box(0); +} +x_184 = lean_array_get_size(x_182); +x_185 = lean_usize_of_nat(x_184); +lean_dec(x_184); +x_186 = lean_usize_sub(x_185, x_168); +x_187 = lean_usize_land(x_166, x_186); +x_188 = lean_array_uget(x_182, x_187); +x_189 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_188); +if (x_189 == 0) +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; +x_190 = lean_unsigned_to_nat(1u); +x_191 = lean_nat_add(x_181, x_190); +lean_dec(x_181); +x_192 = lean_box(0); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_1); +lean_ctor_set(x_193, 1, x_192); +lean_ctor_set(x_193, 2, x_188); +x_194 = lean_array_uset(x_182, x_187, x_193); +x_195 = lean_unsigned_to_nat(4u); +x_196 = lean_nat_mul(x_191, x_195); +x_197 = lean_unsigned_to_nat(3u); +x_198 = lean_nat_div(x_196, x_197); +lean_dec(x_196); +x_199 = lean_array_get_size(x_194); +x_200 = lean_nat_dec_le(x_198, x_199); +lean_dec(x_199); +lean_dec(x_198); +if (x_200 == 0) +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_201 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__2(x_194); +if (lean_is_scalar(x_183)) { + x_202 = lean_alloc_ctor(0, 2, 0); +} else { + x_202 = x_183; +} +lean_ctor_set(x_202, 0, x_191); +lean_ctor_set(x_202, 1, x_201); +if (lean_is_scalar(x_179)) { + x_203 = lean_alloc_ctor(0, 3, 0); +} else { + x_203 = x_179; +} +lean_ctor_set(x_203, 0, x_180); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_178); +x_204 = lean_st_ref_set(x_2, x_203, x_175); +x_205 = lean_ctor_get(x_204, 1); +lean_inc(x_205); +if (lean_is_exclusive(x_204)) { + lean_ctor_release(x_204, 0); + lean_ctor_release(x_204, 1); + x_206 = x_204; +} else { + lean_dec_ref(x_204); + x_206 = lean_box(0); +} +if (lean_is_scalar(x_206)) { + x_207 = lean_alloc_ctor(0, 2, 0); +} else { + x_207 = x_206; +} +lean_ctor_set(x_207, 0, x_192); +lean_ctor_set(x_207, 1, x_205); +return x_207; +} +else +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +if (lean_is_scalar(x_183)) { + x_208 = lean_alloc_ctor(0, 2, 0); +} else { + x_208 = x_183; +} +lean_ctor_set(x_208, 0, x_191); +lean_ctor_set(x_208, 1, x_194); +if (lean_is_scalar(x_179)) { + x_209 = lean_alloc_ctor(0, 3, 0); +} else { + x_209 = x_179; +} +lean_ctor_set(x_209, 0, x_180); +lean_ctor_set(x_209, 1, x_208); +lean_ctor_set(x_209, 2, x_178); +x_210 = lean_st_ref_set(x_2, x_209, x_175); +x_211 = lean_ctor_get(x_210, 1); +lean_inc(x_211); +if (lean_is_exclusive(x_210)) { + lean_ctor_release(x_210, 0); + lean_ctor_release(x_210, 1); + x_212 = x_210; +} else { + lean_dec_ref(x_210); + x_212 = lean_box(0); +} +if (lean_is_scalar(x_212)) { + x_213 = lean_alloc_ctor(0, 2, 0); +} else { + x_213 = x_212; +} +lean_ctor_set(x_213, 0, x_192); +lean_ctor_set(x_213, 1, x_211); +return x_213; +} +} +else +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +lean_dec(x_188); +lean_dec(x_1); +if (lean_is_scalar(x_183)) { + x_214 = lean_alloc_ctor(0, 2, 0); +} else { + x_214 = x_183; +} +lean_ctor_set(x_214, 0, x_181); +lean_ctor_set(x_214, 1, x_182); +if (lean_is_scalar(x_179)) { + x_215 = lean_alloc_ctor(0, 3, 0); +} else { + x_215 = x_179; +} +lean_ctor_set(x_215, 0, x_180); +lean_ctor_set(x_215, 1, x_214); +lean_ctor_set(x_215, 2, x_178); +x_216 = lean_st_ref_set(x_2, x_215, x_175); +x_217 = lean_ctor_get(x_216, 1); +lean_inc(x_217); +if (lean_is_exclusive(x_216)) { + lean_ctor_release(x_216, 0); + lean_ctor_release(x_216, 1); + x_218 = x_216; +} else { + lean_dec_ref(x_216); + x_218 = lean_box(0); +} +x_219 = lean_box(0); +if (lean_is_scalar(x_218)) { + x_220 = lean_alloc_ctor(0, 2, 0); +} else { + x_220 = x_218; +} +lean_ctor_set(x_220, 0, x_219); +lean_ctor_set(x_220, 1, x_217); +return x_220; +} +} +else +{ +lean_object* x_221; lean_object* x_222; +lean_dec(x_1); +x_221 = lean_box(0); +x_222 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_156); +return x_222; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_nat_dec_eq(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_st_ref_get(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 2); +lean_inc(x_10); +lean_dec(x_9); +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; size_t x_23; size_t x_24; size_t x_25; size_t x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; +x_12 = lean_ctor_get(x_8, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_array_get_size(x_13); +x_15 = lean_uint64_of_nat(x_1); +x_16 = 32; +x_17 = lean_uint64_shift_right(x_15, x_16); +x_18 = lean_uint64_xor(x_15, x_17); +x_19 = 16; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = lean_uint64_to_usize(x_21); +x_23 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_24 = 1; +x_25 = lean_usize_sub(x_23, x_24); +x_26 = lean_usize_land(x_22, x_25); +x_27 = lean_array_uget(x_13, x_26); +lean_dec(x_13); +x_28 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_27); +lean_dec(x_27); +x_29 = lean_box(x_28); +lean_ctor_set(x_8, 0, x_29); +return x_8; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; uint64_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; size_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; +x_30 = lean_ctor_get(x_8, 1); +lean_inc(x_30); +lean_dec(x_8); +x_31 = lean_ctor_get(x_10, 1); +lean_inc(x_31); +lean_dec(x_10); +x_32 = lean_array_get_size(x_31); +x_33 = lean_uint64_of_nat(x_1); +x_34 = 32; +x_35 = lean_uint64_shift_right(x_33, x_34); +x_36 = lean_uint64_xor(x_33, x_35); +x_37 = 16; +x_38 = lean_uint64_shift_right(x_36, x_37); +x_39 = lean_uint64_xor(x_36, x_38); +x_40 = lean_uint64_to_usize(x_39); +x_41 = lean_usize_of_nat(x_32); +lean_dec(x_32); +x_42 = 1; +x_43 = lean_usize_sub(x_41, x_42); +x_44 = lean_usize_land(x_40, x_43); +x_45 = lean_array_uget(x_31, x_44); +lean_dec(x_31); +x_46 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_45); +lean_dec(x_45); +x_47 = lean_box(x_46); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_30); +return x_48; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_5); +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = 32; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = 16; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_18 = 1; +x_19 = lean_usize_sub(x_17, x_18); +x_20 = lean_usize_land(x_16, x_19); +x_21 = lean_array_uget(x_2, x_20); +lean_ctor_set(x_3, 2, x_21); +x_22 = lean_array_uset(x_2, x_20, x_3); +x_2 = x_22; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_27 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_24); +x_28 = lean_apply_1(x_1, x_24); +x_29 = lean_unbox_uint64(x_28); +lean_dec(x_28); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_2, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_uset(x_2, x_40, x_42); +x_2 = x_43; +x_3 = x_26; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = lean_uint64_of_nat(x_4); +x_8 = 32; +x_9 = lean_uint64_shift_right(x_7, x_8); +x_10 = lean_uint64_xor(x_7, x_9); +x_11 = 16; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = lean_uint64_to_usize(x_13); +x_15 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_16 = 1; +x_17 = lean_usize_sub(x_15, x_16); +x_18 = lean_usize_land(x_14, x_17); +x_19 = lean_array_uget(x_1, x_18); +lean_ctor_set(x_2, 2, x_19); +x_20 = lean_array_uset(x_1, x_18, x_2); +x_1 = x_20; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; size_t x_33; size_t x_34; size_t x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_22 = lean_ctor_get(x_2, 0); +x_23 = lean_ctor_get(x_2, 1); +x_24 = lean_ctor_get(x_2, 2); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_2); +x_25 = lean_array_get_size(x_1); +x_26 = lean_uint64_of_nat(x_22); +x_27 = 32; +x_28 = lean_uint64_shift_right(x_26, x_27); +x_29 = lean_uint64_xor(x_26, x_28); +x_30 = 16; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = lean_uint64_to_usize(x_32); +x_34 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_35 = 1; +x_36 = lean_usize_sub(x_34, x_35); +x_37 = lean_usize_land(x_33, x_36); +x_38 = lean_array_uget(x_1, x_37); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_23); +lean_ctor_set(x_39, 2, x_38); +x_40 = lean_array_uset(x_1, x_37, x_39); +x_1 = x_40; +x_2 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__3___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__4(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__2(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_st_ref_take(x_2, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 2); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_9, 2); +lean_dec(x_13); +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; uint8_t x_31; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); +x_17 = lean_array_get_size(x_16); +x_18 = lean_uint64_of_nat(x_1); +x_19 = 32; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = 16; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = lean_uint64_to_usize(x_24); +x_26 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_27 = 1; +x_28 = lean_usize_sub(x_26, x_27); +x_29 = lean_usize_land(x_25, x_28); +x_30 = lean_array_uget(x_16, x_29); +x_31 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_32 = lean_unsigned_to_nat(1u); +x_33 = lean_nat_add(x_15, x_32); +lean_dec(x_15); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_30); +x_36 = lean_array_uset(x_16, x_29, x_35); +x_37 = lean_unsigned_to_nat(4u); +x_38 = lean_nat_mul(x_33, x_37); +x_39 = lean_unsigned_to_nat(3u); +x_40 = lean_nat_div(x_38, x_39); +lean_dec(x_38); +x_41 = lean_array_get_size(x_36); +x_42 = lean_nat_dec_le(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_36); +lean_ctor_set(x_10, 1, x_43); +lean_ctor_set(x_10, 0, x_33); +x_44 = lean_st_ref_set(x_2, x_9, x_11); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +lean_object* x_46; +x_46 = lean_ctor_get(x_44, 0); +lean_dec(x_46); +lean_ctor_set(x_44, 0, x_34); +return x_44; +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_34); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +else +{ +lean_object* x_49; uint8_t x_50; +lean_ctor_set(x_10, 1, x_36); +lean_ctor_set(x_10, 0, x_33); +x_49 = lean_st_ref_set(x_2, x_9, x_11); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +lean_ctor_set(x_49, 0, x_34); +return x_49; +} +else +{ +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_49, 1); +lean_inc(x_52); +lean_dec(x_49); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_34); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; uint8_t x_55; +lean_dec(x_30); +lean_dec(x_1); +x_54 = lean_st_ref_set(x_2, x_9, x_11); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_54, 0); +lean_dec(x_56); +x_57 = lean_box(0); +lean_ctor_set(x_54, 0, x_57); +return x_54; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_54, 1); +lean_inc(x_58); +lean_dec(x_54); +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_58); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; size_t x_71; size_t x_72; size_t x_73; size_t x_74; size_t x_75; lean_object* x_76; uint8_t x_77; +x_61 = lean_ctor_get(x_10, 0); +x_62 = lean_ctor_get(x_10, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_10); +x_63 = lean_array_get_size(x_62); +x_64 = lean_uint64_of_nat(x_1); +x_65 = 32; +x_66 = lean_uint64_shift_right(x_64, x_65); +x_67 = lean_uint64_xor(x_64, x_66); +x_68 = 16; +x_69 = lean_uint64_shift_right(x_67, x_68); +x_70 = lean_uint64_xor(x_67, x_69); +x_71 = lean_uint64_to_usize(x_70); +x_72 = lean_usize_of_nat(x_63); +lean_dec(x_63); +x_73 = 1; +x_74 = lean_usize_sub(x_72, x_73); +x_75 = lean_usize_land(x_71, x_74); +x_76 = lean_array_uget(x_62, x_75); +x_77 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_76); +if (x_77 == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_78 = lean_unsigned_to_nat(1u); +x_79 = lean_nat_add(x_61, x_78); +lean_dec(x_61); +x_80 = lean_box(0); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_1); +lean_ctor_set(x_81, 1, x_80); +lean_ctor_set(x_81, 2, x_76); +x_82 = lean_array_uset(x_62, x_75, x_81); +x_83 = lean_unsigned_to_nat(4u); +x_84 = lean_nat_mul(x_79, x_83); +x_85 = lean_unsigned_to_nat(3u); +x_86 = lean_nat_div(x_84, x_85); +lean_dec(x_84); +x_87 = lean_array_get_size(x_82); +x_88 = lean_nat_dec_le(x_86, x_87); +lean_dec(x_87); +lean_dec(x_86); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_89 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_82); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_79); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set(x_9, 2, x_90); +x_91 = lean_st_ref_set(x_2, x_9, x_11); +x_92 = lean_ctor_get(x_91, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_93 = x_91; +} else { + lean_dec_ref(x_91); + x_93 = lean_box(0); +} +if (lean_is_scalar(x_93)) { + x_94 = lean_alloc_ctor(0, 2, 0); +} else { + x_94 = x_93; +} +lean_ctor_set(x_94, 0, x_80); +lean_ctor_set(x_94, 1, x_92); +return x_94; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_79); +lean_ctor_set(x_95, 1, x_82); +lean_ctor_set(x_9, 2, x_95); +x_96 = lean_st_ref_set(x_2, x_9, x_11); +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_98 = x_96; +} else { + lean_dec_ref(x_96); + x_98 = lean_box(0); +} +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_98; +} +lean_ctor_set(x_99, 0, x_80); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_76); +lean_dec(x_1); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_61); +lean_ctor_set(x_100, 1, x_62); +lean_ctor_set(x_9, 2, x_100); +x_101 = lean_st_ref_set(x_2, x_9, x_11); +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_103 = x_101; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); +} +x_104 = lean_box(0); +if (lean_is_scalar(x_103)) { + x_105 = lean_alloc_ctor(0, 2, 0); +} else { + x_105 = x_103; +} +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_102); +return x_105; +} +} +} +else +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint64_t x_112; uint64_t x_113; uint64_t x_114; uint64_t x_115; uint64_t x_116; uint64_t x_117; uint64_t x_118; size_t x_119; size_t x_120; size_t x_121; size_t x_122; size_t x_123; lean_object* x_124; uint8_t x_125; +x_106 = lean_ctor_get(x_9, 0); +x_107 = lean_ctor_get(x_9, 1); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_9); +x_108 = lean_ctor_get(x_10, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_10, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_110 = x_10; +} else { + lean_dec_ref(x_10); + x_110 = lean_box(0); +} +x_111 = lean_array_get_size(x_109); +x_112 = lean_uint64_of_nat(x_1); +x_113 = 32; +x_114 = lean_uint64_shift_right(x_112, x_113); +x_115 = lean_uint64_xor(x_112, x_114); +x_116 = 16; +x_117 = lean_uint64_shift_right(x_115, x_116); +x_118 = lean_uint64_xor(x_115, x_117); +x_119 = lean_uint64_to_usize(x_118); +x_120 = lean_usize_of_nat(x_111); +lean_dec(x_111); +x_121 = 1; +x_122 = lean_usize_sub(x_120, x_121); +x_123 = lean_usize_land(x_119, x_122); +x_124 = lean_array_uget(x_109, x_123); +x_125 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar___spec__1(x_1, x_124); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_126 = lean_unsigned_to_nat(1u); +x_127 = lean_nat_add(x_108, x_126); +lean_dec(x_108); +x_128 = lean_box(0); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_1); +lean_ctor_set(x_129, 1, x_128); +lean_ctor_set(x_129, 2, x_124); +x_130 = lean_array_uset(x_109, x_123, x_129); +x_131 = lean_unsigned_to_nat(4u); +x_132 = lean_nat_mul(x_127, x_131); +x_133 = lean_unsigned_to_nat(3u); +x_134 = lean_nat_div(x_132, x_133); +lean_dec(x_132); +x_135 = lean_array_get_size(x_130); +x_136 = lean_nat_dec_le(x_134, x_135); +lean_dec(x_135); +lean_dec(x_134); +if (x_136 == 0) +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_137 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___spec__1(x_130); +if (lean_is_scalar(x_110)) { + x_138 = lean_alloc_ctor(0, 2, 0); +} else { + x_138 = x_110; +} +lean_ctor_set(x_138, 0, x_127); +lean_ctor_set(x_138, 1, x_137); +x_139 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_139, 0, x_106); +lean_ctor_set(x_139, 1, x_107); +lean_ctor_set(x_139, 2, x_138); +x_140 = lean_st_ref_set(x_2, x_139, x_11); +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_142 = x_140; +} else { + lean_dec_ref(x_140); + x_142 = lean_box(0); +} +if (lean_is_scalar(x_142)) { + x_143 = lean_alloc_ctor(0, 2, 0); +} else { + x_143 = x_142; +} +lean_ctor_set(x_143, 0, x_128); +lean_ctor_set(x_143, 1, x_141); +return x_143; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +if (lean_is_scalar(x_110)) { + x_144 = lean_alloc_ctor(0, 2, 0); +} else { + x_144 = x_110; +} +lean_ctor_set(x_144, 0, x_127); +lean_ctor_set(x_144, 1, x_130); +x_145 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_145, 0, x_106); +lean_ctor_set(x_145, 1, x_107); +lean_ctor_set(x_145, 2, x_144); +x_146 = lean_st_ref_set(x_2, x_145, x_11); +x_147 = lean_ctor_get(x_146, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_148 = x_146; +} else { + lean_dec_ref(x_146); + x_148 = lean_box(0); +} +if (lean_is_scalar(x_148)) { + x_149 = lean_alloc_ctor(0, 2, 0); +} else { + x_149 = x_148; +} +lean_ctor_set(x_149, 0, x_128); +lean_ctor_set(x_149, 1, x_147); +return x_149; +} +} +else +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +lean_dec(x_124); +lean_dec(x_1); +if (lean_is_scalar(x_110)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_110; +} +lean_ctor_set(x_150, 0, x_108); +lean_ctor_set(x_150, 1, x_109); +x_151 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_151, 0, x_106); +lean_ctor_set(x_151, 1, x_107); +lean_ctor_set(x_151, 2, x_150); +x_152 = lean_st_ref_set(x_2, x_151, x_11); +x_153 = lean_ctor_get(x_152, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_152)) { + lean_ctor_release(x_152, 0); + lean_ctor_release(x_152, 1); + x_154 = x_152; +} else { + lean_dec_ref(x_152); + x_154 = lean_box(0); +} +x_155 = lean_box(0); +if (lean_is_scalar(x_154)) { + x_156 = lean_alloc_ctor(0, 2, 0); +} else { + x_156 = x_154; +} +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_153); +return x_156; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_Lean_Expr_isApp(x_1); +if (x_2 == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = l_Lean_Expr_getAppFn(x_1); +switch (lean_obj_tag(x_4)) { +case 1: +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_4); +return x_5; +} +case 4: +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +x_8 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_7, x_6); +lean_dec(x_6); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_4); +return x_9; +} +else +{ +lean_object* x_10; +lean_dec(x_4); +x_10 = lean_box(0); +return x_10; +} +} +default: +{ +lean_object* x_11; +lean_dec(x_4); +x_11 = lean_box(0); +return x_11; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_2, x_1); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_3); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; uint8_t x_22; lean_object* x_23; lean_object* x_38; +x_11 = lean_array_uget(x_3, x_2); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_array_uset(x_3, x_2, x_12); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_11); +x_38 = l_Lean_Meta_isProp(x_11, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; uint8_t x_40; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_unbox(x_39); +lean_dec(x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +lean_dec(x_38); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_11); +x_42 = l_Lean_Meta_isTypeFormer(x_11, x_4, x_5, x_6, x_7, x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; uint8_t x_44; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_unbox(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_43); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_11); +x_46 = l_Lean_Meta_isProof(x_11, x_4, x_5, x_6, x_7, x_45); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_unbox(x_47); +lean_dec(x_47); +x_22 = x_49; +x_23 = x_48; +goto block_37; +} +else +{ +uint8_t x_50; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_50 = !lean_is_exclusive(x_46); +if (x_50 == 0) +{ +return x_46; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_46, 0); +x_52 = lean_ctor_get(x_46, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_46); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; uint8_t x_55; +x_54 = lean_ctor_get(x_42, 1); +lean_inc(x_54); +lean_dec(x_42); +x_55 = lean_unbox(x_43); +lean_dec(x_43); +x_22 = x_55; +x_23 = x_54; +goto block_37; +} +} +else +{ +uint8_t x_56; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_56 = !lean_is_exclusive(x_42); +if (x_56 == 0) +{ +return x_42; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_42, 0); +x_58 = lean_ctor_get(x_42, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_42); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +else +{ +lean_object* x_60; size_t x_61; size_t x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_11); +x_60 = lean_ctor_get(x_38, 1); +lean_inc(x_60); +lean_dec(x_38); +x_61 = 1; +x_62 = lean_usize_add(x_2, x_61); +x_63 = 0; +x_64 = lean_box(x_63); +x_65 = lean_array_uset(x_13, x_2, x_64); +x_2 = x_62; +x_3 = x_65; +x_8 = x_60; +goto _start; +} +} +else +{ +uint8_t x_67; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_67 = !lean_is_exclusive(x_38); +if (x_67 == 0) +{ +return x_38; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_38, 0); +x_69 = lean_ctor_get(x_38, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_38); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +block_21: +{ +size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_16 = 1; +x_17 = lean_usize_add(x_2, x_16); +x_18 = lean_box(x_14); +x_19 = lean_array_uset(x_13, x_2, x_18); +x_2 = x_17; +x_3 = x_19; +x_8 = x_15; +goto _start; +} +block_37: +{ +if (x_22 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = l_Lean_Expr_fvarId_x21(x_11); +lean_dec(x_11); +lean_inc(x_4); +x_25 = l_Lean_FVarId_getDecl(x_24, x_4, x_5, x_6, x_7, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_LocalDecl_binderInfo(x_26); +lean_dec(x_26); +x_29 = lean_box(x_28); +if (lean_obj_tag(x_29) == 3) +{ +uint8_t x_30; +x_30 = 1; +x_14 = x_30; +x_15 = x_27; +goto block_21; +} +else +{ +uint8_t x_31; +lean_dec(x_29); +x_31 = 0; +x_14 = x_31; +x_15 = x_27; +goto block_21; +} +} +else +{ +uint8_t x_32; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) +{ +return x_25; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_25, 0); +x_34 = lean_ctor_get(x_25, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_25); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_11); +x_36 = 1; +x_14 = x_36; +x_15 = x_23; +goto block_21; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_array_size(x_1); +x_9 = 0; +x_10 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1(x_8, x_9, x_1, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1___boxed), 7, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_2); +x_12 = l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1; +x_13 = 0; +x_14 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_9, x_11, x_12, x_13, x_3, x_4, x_5, x_6, x_10); +return x_14; +} +else +{ +uint8_t x_15; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_8); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_10 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_11 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___spec__1(x_9, x_10, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = l_Lean_Expr_hasLooseBVars(x_1); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_10 = l_Lean_Expr_hasMVar(x_1); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Meta_Grind_mkGroundPattern(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_1); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_8); +return x_14; +} +} +else +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_foundBVar(x_15, x_3, x_4, x_5, x_6, x_7, x_8); +if (x_2 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_15, x_3, x_4, x_5, x_6, x_7, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_1); +return x_18; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_1); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_dec(x_16); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveBVar(x_15, x_3, x_4, x_5, x_6, x_7, x_25); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +lean_ctor_set(x_26, 0, x_1); +return x_26; +} +else +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_dec(x_26); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +else +{ +uint8_t x_31; +lean_dec(x_15); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_16); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_16, 0); +lean_dec(x_32); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +lean_ctor_set(x_16, 0, x_33); +return x_16; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_16, 1); +lean_inc(x_34); +lean_dec(x_16); +x_35 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +return x_36; +} +} +} +} +else +{ +if (x_2 == 0) +{ +lean_object* x_37; +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_38 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_8); +return x_39; +} +else +{ +uint8_t x_40; lean_object* x_41; +lean_dec(x_37); +x_40 = 0; +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_40, x_3, x_4, x_5, x_6, x_7, x_8); +return x_41; +} +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_42 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_8); +return x_43; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 5); +x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set_tag(x_9, 1); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +lean_inc(x_8); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_5, 1); +x_17 = lean_nat_dec_lt(x_7, x_16); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_6); +lean_ctor_set(x_18, 1, x_15); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_array_fget(x_6, x_7); +x_20 = lean_array_get_size(x_3); +x_21 = lean_nat_dec_lt(x_7, x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; lean_object* x_23; +x_22 = 0; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(x_19, x_22, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_array_fset(x_6, x_7, x_24); +x_27 = lean_ctor_get(x_5, 2); +x_28 = lean_nat_add(x_7, x_27); +lean_dec(x_7); +x_6 = x_26; +x_7 = x_28; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_15 = x_25; +goto _start; +} +else +{ +uint8_t x_30; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) +{ +return x_23; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_23, 0); +x_32 = lean_ctor_get(x_23, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_23); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +else +{ +lean_object* x_34; uint8_t x_35; lean_object* x_36; +x_34 = lean_array_fget(x_3, x_7); +x_35 = lean_unbox(x_34); +lean_dec(x_34); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_36 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(x_19, x_35, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_array_fset(x_6, x_7, x_37); +x_40 = lean_ctor_get(x_5, 2); +x_41 = lean_nat_add(x_7, x_40); +lean_dec(x_7); +x_6 = x_39; +x_7 = x_41; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_15 = x_38; +goto _start; +} +else +{ +uint8_t x_43; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +x_43 = !lean_is_exclusive(x_36); +if (x_43 == 0) +{ +return x_36; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_36, 0); +x_45 = lean_ctor_get(x_36, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_36); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +} +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 5); +x_9 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set_tag(x_9, 1); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +lean_inc(x_8); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern, (non-forbidden) application expected", 53, 53); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("f.isConst || f.isFVar\n ", 24, 24); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__4; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.NormalizePattern.go", 83, 83); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__6; +x_3 = lean_unsigned_to_nat(266u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__5; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFn_x3f(x_1); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = l_Lean_indentExpr(x_1); +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__2; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_14, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; uint8_t x_50; +x_16 = lean_ctor_get(x_9, 0); +lean_inc(x_16); +lean_dec(x_9); +x_50 = l_Lean_Expr_isConst(x_16); +if (x_50 == 0) +{ +uint8_t x_51; +x_51 = l_Lean_Expr_isFVar(x_16); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; +lean_dec(x_16); +lean_dec(x_1); +x_52 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__7; +x_53 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3(x_52, x_3, x_4, x_5, x_6, x_7, x_8); +return x_53; +} +else +{ +lean_object* x_54; +x_54 = lean_box(0); +x_17 = x_54; +goto block_49; +} +} +else +{ +lean_object* x_55; +x_55 = lean_box(0); +x_17 = x_55; +goto block_49; +} +block_49: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_17); +lean_inc(x_16); +x_18 = l_Lean_Expr_toHeadIndex(x_16); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_saveSymbol(x_18, x_3, x_4, x_5, x_6, x_7, x_8); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_unsigned_to_nat(0u); +x_22 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_21); +x_23 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_22); +x_24 = lean_mk_array(x_22, x_23); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_sub(x_22, x_25); +lean_dec(x_22); +lean_inc(x_1); +x_27 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_24, x_26); +x_28 = lean_array_get_size(x_27); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_28); +lean_inc(x_16); +x_29 = l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask(x_16, x_28, x_4, x_5, x_6, x_7, x_20); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +lean_inc(x_28); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_21); +lean_ctor_set(x_32, 1, x_28); +lean_ctor_set(x_32, 2, x_25); +x_33 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_1, x_28, x_30, x_32, x_32, x_27, x_21, lean_box(0), lean_box(0), x_3, x_4, x_5, x_6, x_7, x_31); +lean_dec(x_32); +lean_dec(x_30); +lean_dec(x_28); +lean_dec(x_1); +if (lean_obj_tag(x_33) == 0) +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +x_36 = l_Lean_mkAppN(x_16, x_35); +lean_dec(x_35); +lean_ctor_set(x_33, 0, x_36); +return x_33; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_33, 0); +x_38 = lean_ctor_get(x_33, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_33); +x_39 = l_Lean_mkAppN(x_16, x_37); +lean_dec(x_37); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +return x_40; +} +} +else +{ +uint8_t x_41; +lean_dec(x_16); +x_41 = !lean_is_exclusive(x_33); +if (x_41 == 0) +{ +return x_33; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_33, 0); +x_43 = lean_ctor_get(x_33, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_33); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +else +{ +uint8_t x_45; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_16); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_45 = !lean_is_exclusive(x_29); +if (x_45 == 0) +{ +return x_29; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_29, 0); +x_47 = lean_ctor_get(x_29, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_29); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_1); +x_9 = l_Lean_Meta_Grind_isOffsetPattern_x3f(x_1); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_box(0); +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_dec(x_1); +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = 0; +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(x_13, x_15, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_16, 0); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_20 = lean_expr_eqv(x_18, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = l_Lean_Meta_Grind_mkOffsetPattern(x_18, x_14); +lean_ctor_set(x_16, 0, x_21); +return x_16; +} +else +{ +lean_dec(x_18); +lean_dec(x_14); +lean_ctor_set(x_16, 0, x_19); +return x_16; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_16); +x_24 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_25 = lean_expr_eqv(x_22, x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = l_Lean_Meta_Grind_mkOffsetPattern(x_22, x_14); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +return x_27; +} +else +{ +lean_object* x_28; +lean_dec(x_22); +lean_dec(x_14); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_24); +lean_ctor_set(x_28, 1, x_23); +return x_28; +} +} +} +else +{ +uint8_t x_29; +lean_dec(x_14); +x_29 = !lean_is_exclusive(x_16); +if (x_29 == 0) +{ +return x_16; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_16, 0); +x_31 = lean_ctor_get(x_16, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_16); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern, it does not have pattern variables", 51, 51); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (x_2 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = l_Lean_Expr_hasLooseBVars(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +lean_dec(x_1); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___closed__2; +x_13 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_12, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +return x_13; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_13); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_box(0); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go_goArg(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_9 = l_List_reverse___rarg(x_2); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_ctor_get(x_1, 1); +x_14 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_12, x_14, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_16); +{ +lean_object* _tmp_0 = x_13; +lean_object* _tmp_1 = x_1; +lean_object* _tmp_7 = x_17; +x_1 = _tmp_0; +x_2 = _tmp_1; +x_8 = _tmp_7; +} +goto _start; +} +else +{ +uint8_t x_19; +lean_free_object(x_1); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) +{ +return x_15; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 0); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_1); +x_25 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_23, x_25, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_2); +x_1 = x_24; +x_2 = x_29; +x_8 = x_28; +goto _start; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_24); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_31 = lean_ctor_get(x_26, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_33 = x_26; +} else { + lean_dec_ref(x_26); + x_33 = lean_box(0); +} +if (lean_is_scalar(x_33)) { + x_34 = lean_alloc_ctor(1, 2, 0); +} else { + x_34 = x_33; +} +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__2; +x_3 = lean_mk_array(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_main(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = lean_box(0); +x_8 = l_Lean_Meta_Grind_NormalizePattern_main___closed__5; +x_9 = lean_st_mk_ref(x_8, x_6); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +x_13 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_11, x_2, x_3, x_4, x_5, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_st_ref_get(x_11, x_15); +lean_dec(x_11); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_array_to_list(x_19); +x_21 = lean_ctor_get(x_18, 2); +lean_inc(x_21); +lean_dec(x_18); +lean_ctor_set(x_9, 1, x_21); +lean_ctor_set(x_9, 0, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_14); +lean_ctor_set(x_22, 1, x_9); +lean_ctor_set(x_16, 0, x_22); +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +x_26 = lean_array_to_list(x_25); +x_27 = lean_ctor_get(x_23, 2); +lean_inc(x_27); +lean_dec(x_23); +lean_ctor_set(x_9, 1, x_27); +lean_ctor_set(x_9, 0, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_14); +lean_ctor_set(x_28, 1, x_9); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_24); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_free_object(x_9); +lean_dec(x_11); +x_30 = !lean_is_exclusive(x_13); +if (x_30 == 0) +{ +return x_13; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_13, 0); +x_32 = lean_ctor_get(x_13, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_13); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_9, 0); +x_35 = lean_ctor_get(x_9, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_9); +lean_inc(x_34); +x_36 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_34, x_2, x_3, x_4, x_5, x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_st_ref_get(x_34, x_38); +lean_dec(x_34); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_42 = x_39; +} else { + lean_dec_ref(x_39); + x_42 = lean_box(0); +} +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +x_44 = lean_array_to_list(x_43); +x_45 = lean_ctor_get(x_40, 2); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_37); +lean_ctor_set(x_47, 1, x_46); +if (lean_is_scalar(x_42)) { + x_48 = lean_alloc_ctor(0, 2, 0); +} else { + x_48 = x_42; +} +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_41); +return x_48; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_34); +x_49 = lean_ctor_get(x_36, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_36, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_51 = x_36; +} else { + lean_dec_ref(x_36); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_normalizePattern(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = 0; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_8, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_array_uget(x_4, x_5); +x_9 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_8); +if (lean_obj_tag(x_9) == 0) +{ +size_t x_10; size_t x_11; +lean_dec(x_8); +x_10 = 1; +x_11 = lean_usize_add(x_5, x_10); +x_5 = x_11; +goto _start; +} +else +{ +lean_object* x_13; +lean_dec(x_9); +x_13 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_8); +lean_dec(x_8); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +else +{ +size_t x_15; size_t x_16; +lean_dec(x_13); +x_15 = 1; +x_16 = lean_usize_add(x_5, x_15); +x_5 = x_16; +goto _start; +} +} +} +else +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_4, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_array_uget(x_3, x_4); +x_8 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_7); +if (lean_obj_tag(x_8) == 0) +{ +size_t x_9; size_t x_10; +lean_dec(x_7); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_12; +lean_dec(x_8); +x_12 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_7); +lean_dec(x_7); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = 1; +return x_13; +} +else +{ +size_t x_14; size_t x_15; +lean_dec(x_12); +x_14 = 1; +x_15 = lean_usize_add(x_4, x_14); +x_4 = x_15; +goto _start; +} +} +} +else +{ +uint8_t x_17; +x_17 = 0; +return x_17; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_3 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_1); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1; +x_5 = l_Lean_CollectFVars_main(x_3, x_4); +x_6 = lean_ctor_get(x_5, 2); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_7); +lean_dec(x_6); +x_10 = 1; +return x_10; +} +else +{ +size_t x_11; size_t x_12; uint8_t x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_13 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_6, x_11, x_12); +lean_dec(x_6); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +else +{ +uint8_t x_15; +x_15 = 0; +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; uint8_t x_9; lean_object* x_10; +x_7 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_8 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_9 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(x_1, x_2, x_3, x_4, x_7, x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_10 = lean_box(x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; uint8_t x_8; lean_object* x_9; +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_8 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_3, x_6, x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_9); +x_17 = lean_array_uget(x_6, x_8); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_18 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +x_22 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_20); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; +lean_ctor_set(x_18, 0, x_23); +return x_18; +} +else +{ +size_t x_24; size_t x_25; +lean_free_object(x_18); +x_24 = 1; +x_25 = lean_usize_add(x_8, x_24); +lean_inc(x_5); +{ +size_t _tmp_7 = x_25; +lean_object* _tmp_8 = x_5; +lean_object* _tmp_13 = x_21; +x_8 = _tmp_7; +x_9 = _tmp_8; +x_14 = _tmp_13; +} +goto _start; +} +} +else +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_18); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_30 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +return x_31; +} +else +{ +size_t x_32; size_t x_33; +x_32 = 1; +x_33 = lean_usize_add(x_8, x_32); +lean_inc(x_5); +{ +size_t _tmp_7 = x_33; +lean_object* _tmp_8 = x_5; +lean_object* _tmp_13 = x_28; +x_8 = _tmp_7; +x_9 = _tmp_8; +x_14 = _tmp_13; +} +goto _start; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +x_35 = !lean_is_exclusive(x_18); +if (x_35 == 0) +{ +return x_18; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_18, 0); +x_37 = lean_ctor_get(x_18, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_18); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("semiOutParam", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("outParam", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_27; +x_17 = lean_array_uget(x_6, x_8); +x_27 = !lean_is_exclusive(x_9); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_28 = lean_ctor_get(x_9, 1); +x_29 = lean_ctor_get(x_9, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_28, 2); +lean_inc(x_32); +x_33 = lean_nat_dec_lt(x_31, x_32); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_17); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_9); +x_18 = x_34; +x_19 = x_14; +goto block_26; +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_36 = lean_ctor_get(x_28, 2); +lean_dec(x_36); +x_37 = lean_ctor_get(x_28, 1); +lean_dec(x_37); +x_38 = lean_ctor_get(x_28, 0); +lean_dec(x_38); +x_39 = lean_array_fget(x_30, x_31); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_add(x_31, x_40); +lean_dec(x_31); +lean_ctor_set(x_28, 1, x_41); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_42 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_46 = l_Lean_Expr_isAppOf(x_43, x_45); +if (x_46 == 0) +{ +lean_object* x_47; uint8_t x_48; +x_47 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_48 = l_Lean_Expr_isAppOf(x_43, x_47); +lean_dec(x_43); +if (x_48 == 0) +{ +uint8_t x_49; +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_39); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +lean_ctor_set(x_9, 0, x_50); +x_51 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_51, 0, x_9); +x_18 = x_51; +x_19 = x_44; +goto block_26; +} +else +{ +lean_object* x_52; +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_9); +x_18 = x_52; +x_19 = x_44; +goto block_26; +} +} +else +{ +lean_object* x_53; +lean_dec(x_39); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_9); +x_18 = x_53; +x_19 = x_44; +goto block_26; +} +} +else +{ +lean_object* x_54; +lean_dec(x_43); +lean_dec(x_39); +lean_inc(x_4); +lean_ctor_set(x_9, 0, x_4); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_9); +x_18 = x_54; +x_19 = x_44; +goto block_26; +} +} +else +{ +uint8_t x_55; +lean_dec(x_28); +lean_dec(x_39); +lean_free_object(x_9); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_55 = !lean_is_exclusive(x_42); +if (x_55 == 0) +{ +return x_42; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_42, 0); +x_57 = lean_ctor_get(x_42, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_42); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_28); +x_59 = lean_array_fget(x_30, x_31); +x_60 = lean_unsigned_to_nat(1u); +x_61 = lean_nat_add(x_31, x_60); +lean_dec(x_31); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_30); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_32); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_63 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_67 = l_Lean_Expr_isAppOf(x_64, x_66); +if (x_67 == 0) +{ +lean_object* x_68; uint8_t x_69; +x_68 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_69 = l_Lean_Expr_isAppOf(x_64, x_68); +lean_dec(x_64); +if (x_69 == 0) +{ +uint8_t x_70; +x_70 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_59); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +x_71 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_71); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_9); +x_18 = x_72; +x_19 = x_65; +goto block_26; +} +else +{ +lean_object* x_73; +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_73 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_73, 0, x_9); +x_18 = x_73; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_74; +lean_dec(x_59); +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_9); +x_18 = x_74; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_75; +lean_dec(x_64); +lean_dec(x_59); +lean_inc(x_4); +lean_ctor_set(x_9, 1, x_62); +lean_ctor_set(x_9, 0, x_4); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_9); +x_18 = x_75; +x_19 = x_65; +goto block_26; +} +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_62); +lean_dec(x_59); +lean_free_object(x_9); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_76 = lean_ctor_get(x_63, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_63, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_78 = x_63; +} else { + lean_dec_ref(x_63); + x_78 = lean_box(0); +} +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); +} else { + x_79 = x_78; +} +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_80 = lean_ctor_get(x_9, 1); +lean_inc(x_80); +lean_dec(x_9); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +x_83 = lean_ctor_get(x_80, 2); +lean_inc(x_83); +x_84 = lean_nat_dec_lt(x_82, x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; +lean_dec(x_83); +lean_dec(x_82); +lean_dec(x_81); +lean_dec(x_17); +lean_inc(x_4); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_4); +lean_ctor_set(x_85, 1, x_80); +x_86 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_86, 0, x_85); +x_18 = x_86; +x_19 = x_14; +goto block_26; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + lean_ctor_release(x_80, 2); + x_87 = x_80; +} else { + lean_dec_ref(x_80); + x_87 = lean_box(0); +} +x_88 = lean_array_fget(x_81, x_82); +x_89 = lean_unsigned_to_nat(1u); +x_90 = lean_nat_add(x_82, x_89); +lean_dec(x_82); +if (lean_is_scalar(x_87)) { + x_91 = lean_alloc_ctor(0, 3, 0); +} else { + x_91 = x_87; +} +lean_ctor_set(x_91, 0, x_81); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_83); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_92 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +x_95 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; +x_96 = l_Lean_Expr_isAppOf(x_93, x_95); +if (x_96 == 0) +{ +lean_object* x_97; uint8_t x_98; +x_97 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; +x_98 = l_Lean_Expr_isAppOf(x_93, x_97); +lean_dec(x_93); +if (x_98 == 0) +{ +uint8_t x_99; +x_99 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_88); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_91); +x_102 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_102, 0, x_101); +x_18 = x_102; +x_19 = x_94; +goto block_26; +} +else +{ +lean_object* x_103; lean_object* x_104; +lean_inc(x_4); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_4); +lean_ctor_set(x_103, 1, x_91); +x_104 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_104, 0, x_103); +x_18 = x_104; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_105; lean_object* x_106; +lean_dec(x_88); +lean_inc(x_4); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_4); +lean_ctor_set(x_105, 1, x_91); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_18 = x_106; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_dec(x_93); +lean_dec(x_88); +lean_inc(x_4); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_4); +lean_ctor_set(x_107, 1, x_91); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_107); +x_18 = x_108; +x_19 = x_94; +goto block_26; +} +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_dec(x_91); +lean_dec(x_88); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_109 = lean_ctor_get(x_92, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_92, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_111 = x_92; +} else { + lean_dec_ref(x_92); + x_111 = lean_box(0); +} +if (lean_is_scalar(x_111)) { + x_112 = lean_alloc_ctor(1, 2, 0); +} else { + x_112 = x_111; +} +lean_ctor_set(x_112, 0, x_109); +lean_ctor_set(x_112, 1, x_110); +return x_112; +} +} +} +block_26: +{ +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +else +{ +lean_object* x_22; size_t x_23; size_t x_24; +x_22 = lean_ctor_get(x_18, 0); +lean_inc(x_22); +lean_dec(x_18); +x_23 = 1; +x_24 = lean_usize_add(x_8, x_23); +x_8 = x_24; +x_9 = x_22; +x_14 = x_19; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_7 = 1; +x_8 = lean_box(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; +x_14 = lean_array_get_size(x_1); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_toSubarray___rarg(x_1, x_15, x_14); +lean_inc(x_2); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, x_16); +x_18 = lean_array_size(x_7); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_3, x_4, x_5, x_2, x_7, x_7, x_18, x_6, x_17, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1; +x_24 = lean_box(0); +x_25 = lean_apply_6(x_23, x_24, x_9, x_10, x_11, x_12, x_22); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_19, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_21, 0); +lean_inc(x_28); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_28); +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_dec(x_19); +x_30 = lean_ctor_get(x_21, 0); +lean_inc(x_30); +lean_dec(x_21); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +} +else +{ +uint8_t x_32; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_32 = !lean_is_exclusive(x_19); +if (x_32 == 0) +{ +return x_19; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_19, 0); +x_34 = lean_ctor_get(x_19, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_19); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_15 = lean_infer_type(x_1, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unsigned_to_nat(0u); +x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_box_usize(x_8); +x_22 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed), 13, 6); +lean_closure_set(x_22, 0, x_3); +lean_closure_set(x_22, 1, x_4); +lean_closure_set(x_22, 2, x_5); +lean_closure_set(x_22, 3, x_6); +lean_closure_set(x_22, 4, x_7); +lean_closure_set(x_22, 5, x_21); +x_23 = 0; +x_24 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_16, x_20, x_22, x_23, x_10, x_11, x_12, x_13, x_17); +return x_24; +} +else +{ +uint8_t x_25; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_25 = !lean_is_exclusive(x_15); +if (x_25 == 0) +{ +return x_15; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_5) == 5) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_5, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_5, 1); +lean_inc(x_14); +lean_dec(x_5); +x_15 = lean_array_set(x_6, x_7, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_7, x_16); +lean_dec(x_7); +x_5 = x_13; +x_6 = x_15; +x_7 = x_17; +goto _start; +} +else +{ +lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_7); +x_19 = lean_box(0); +x_20 = lean_box(0); +x_21 = lean_array_size(x_3); +x_22 = 0; +x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_24 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_19, x_23, x_3, x_21, x_22, x_23, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec(x_25); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_box(0); +x_29 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_5, x_4, x_6, x_20, x_1, x_2, x_19, x_22, x_28, x_8, x_9, x_10, x_11, x_27); +return x_29; +} +else +{ +uint8_t x_30; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_24); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_24, 0); +lean_dec(x_31); +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +lean_ctor_set(x_24, 0, x_32); +return x_24; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_24, 1); +lean_inc(x_33); +lean_dec(x_24); +x_34 = lean_ctor_get(x_26, 0); +lean_inc(x_34); +lean_dec(x_26); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_24); +if (x_36 == 0) +{ +return x_24; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_24, 0); +x_38 = lean_ctor_get(x_24, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_24); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_unsigned_to_nat(0u); +x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_10); +x_12 = l_Lean_Meta_Grind_ppPattern___closed__5; +lean_inc(x_11); +x_13 = lean_mk_array(x_11, x_12); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_11, x_14); +lean_dec(x_11); +lean_inc(x_4); +x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_4, x_13, x_15, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1___boxed), 9, 2); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_2); +x_10 = 0; +x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_3, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; lean_object* x_15; +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(x_1, x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; lean_object* x_16; +x_15 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_9); +lean_dec(x_2); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +lean_dec(x_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_3); +return x_10; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instInhabitedMetaM___boxed), 5, 1); +lean_closure_set(x_1, 0, lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; +x_8 = lean_panic_fn(x_7, x_1); +x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_inc(x_1); +return x_1; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 2); +x_5 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_4); +lean_inc(x_3); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Lean_Expr_fvarId_x21(x_5); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_Expr_fvarId_x21(x_9); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_4); +x_6 = lean_box(0); +x_7 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_3, x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_6; +x_6 = l_List_reverse___rarg(x_5); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_4); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +x_10 = lean_nat_sub(x_1, x_8); +lean_dec(x_8); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_10, x_11); +lean_dec(x_10); +x_13 = lean_nat_dec_lt(x_12, x_3); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_12); +x_14 = l_Lean_instInhabitedExpr; +x_15 = l_outOfBounds___rarg(x_14); +x_16 = l_Lean_Expr_fvarId_x21(x_15); +lean_dec(x_15); +lean_ctor_set(x_4, 1, x_5); +lean_ctor_set(x_4, 0, x_16); +{ +lean_object* _tmp_3 = x_9; +lean_object* _tmp_4 = x_4; +x_4 = _tmp_3; +x_5 = _tmp_4; +} +goto _start; +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_array_fget(x_2, x_12); +lean_dec(x_12); +x_19 = l_Lean_Expr_fvarId_x21(x_18); +lean_dec(x_18); +lean_ctor_set(x_4, 1, x_5); +lean_ctor_set(x_4, 0, x_19); +{ +lean_object* _tmp_3 = x_9; +lean_object* _tmp_4 = x_4; +x_4 = _tmp_3; +x_5 = _tmp_4; +} +goto _start; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_21 = lean_ctor_get(x_4, 0); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_4); +x_23 = lean_nat_sub(x_1, x_21); +lean_dec(x_21); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_23, x_24); +lean_dec(x_23); +x_26 = lean_nat_dec_lt(x_25, x_3); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_25); +x_27 = l_Lean_instInhabitedExpr; +x_28 = l_outOfBounds___rarg(x_27); +x_29 = l_Lean_Expr_fvarId_x21(x_28); +lean_dec(x_28); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_5); +x_4 = x_22; +x_5 = x_30; +goto _start; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_array_fget(x_2, x_25); +lean_dec(x_25); +x_33 = l_Lean_Expr_fvarId_x21(x_32); +lean_dec(x_32); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_5); +x_4 = x_22; +x_5 = x_34; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_15; +lean_dec(x_10); +lean_dec(x_1); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_7, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_7, 1); +lean_inc(x_17); +lean_dec(x_7); +lean_inc(x_10); +x_18 = l_Lean_FVarId_getType(x_16, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_box(0); +lean_inc(x_1); +x_22 = lean_array_mk(x_1); +x_23 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_24 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_22); +x_25 = l_Lean_CollectFVars_main(x_19, x_24); +x_26 = lean_ctor_get(x_25, 2); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_array_get_size(x_26); +x_28 = lean_unsigned_to_nat(0u); +x_29 = lean_nat_dec_lt(x_28, x_27); +if (x_29 == 0) +{ +lean_dec(x_27); +lean_dec(x_26); +x_7 = x_17; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +else +{ +uint8_t x_31; +x_31 = lean_nat_dec_le(x_27, x_27); +if (x_31 == 0) +{ +lean_dec(x_27); +lean_dec(x_26); +x_7 = x_17; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +else +{ +size_t x_33; size_t x_34; lean_object* x_35; +x_33 = 0; +x_34 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_35 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_3, x_26, x_33, x_34, x_8); +lean_dec(x_26); +x_7 = x_17; +x_8 = x_35; +x_9 = lean_box(0); +x_14 = x_20; +goto _start; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_18); +if (x_37 == 0) +{ +return x_18; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_18, 0); +x_39 = lean_ctor_get(x_18, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_18); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_14; +lean_dec(x_9); +lean_dec(x_1); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_6, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_6, 1); +lean_inc(x_16); +lean_dec(x_6); +lean_inc(x_9); +x_17 = l_Lean_FVarId_getType(x_15, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +lean_inc(x_1); +x_21 = lean_array_mk(x_1); +x_22 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_23 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +lean_ctor_set(x_23, 2, x_21); +x_24 = l_Lean_CollectFVars_main(x_18, x_23); +x_25 = lean_ctor_get(x_24, 2); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_array_get_size(x_25); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_nat_dec_lt(x_27, x_26); +if (x_28 == 0) +{ +lean_dec(x_26); +lean_dec(x_25); +x_6 = x_16; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +else +{ +uint8_t x_30; +x_30 = lean_nat_dec_le(x_26, x_26); +if (x_30 == 0) +{ +lean_dec(x_26); +lean_dec(x_25); +x_6 = x_16; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +else +{ +size_t x_32; size_t x_33; lean_object* x_34; +x_32 = 0; +x_33 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_34 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_2, x_25, x_32, x_33, x_7); +lean_dec(x_25); +x_6 = x_16; +x_7 = x_34; +x_8 = lean_box(0); +x_13 = x_19; +goto _start; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_17); +if (x_36 == 0) +{ +return x_17; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_17, 0); +x_38 = lean_ctor_get(x_17, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_17); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_array_uget(x_3, x_4); +x_9 = 1; +x_10 = lean_usize_add(x_4, x_9); +x_11 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_8); +x_4 = x_10; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); +x_4 = x_10; +x_6 = x_14; +goto _start; +} +} +else +{ +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_10 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_1, x_7); +if (lean_obj_tag(x_10) == 0) +{ +lean_dec(x_7); +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); +x_3 = x_9; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_8, x_7); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_25; +x_17 = lean_array_uget(x_6, x_8); +x_25 = !lean_is_exclusive(x_9); +if (x_25 == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_9, 1); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_9, 0); +x_29 = lean_ctor_get(x_26, 0); +x_30 = lean_ctor_get(x_26, 1); +x_31 = l_Lean_Expr_fvarId_x21(x_17); +x_32 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_30, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_33 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_28, x_31); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_34); +x_37 = l_Lean_Meta_isProp(x_34, x_10, x_11, x_12, x_13, x_35); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; uint8_t x_39; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_unbox(x_38); +lean_dec(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +lean_inc(x_10); +lean_inc(x_31); +x_41 = l_Lean_FVarId_getDecl(x_31, x_10, x_11, x_12, x_13, x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Lean_LocalDecl_binderInfo(x_42); +lean_dec(x_42); +x_45 = lean_box(x_44); +if (lean_obj_tag(x_45) == 3) +{ +lean_object* x_46; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_34); +lean_inc(x_28); +lean_inc(x_4); +x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_28, x_34, x_10, x_11, x_12, x_13, x_43); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; uint8_t x_48; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_unbox(x_47); +lean_dec(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; +lean_dec(x_34); +lean_dec(x_31); +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +lean_dec(x_46); +x_50 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_50, 0, x_9); +x_18 = x_50; +x_19 = x_49; +goto block_24; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; +lean_dec(x_29); +x_51 = lean_ctor_get(x_46, 1); +lean_inc(x_51); +lean_dec(x_46); +x_52 = lean_box(0); +lean_inc(x_31); +x_53 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_52); +x_54 = lean_box(0); +lean_inc(x_2); +x_55 = lean_array_mk(x_2); +x_56 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +lean_ctor_set(x_57, 2, x_55); +x_58 = l_Lean_CollectFVars_main(x_34, x_57); +x_59 = lean_ctor_get(x_58, 2); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_array_get_size(x_59); +x_61 = lean_unsigned_to_nat(0u); +x_62 = lean_nat_dec_lt(x_61, x_60); +x_63 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_52); +if (x_62 == 0) +{ +uint8_t x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_60); +lean_dec(x_59); +x_64 = 1; +x_65 = lean_box(x_64); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_65); +lean_ctor_set(x_9, 0, x_53); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_9); +x_18 = x_66; +x_19 = x_51; +goto block_24; +} +else +{ +uint8_t x_67; +x_67 = lean_nat_dec_le(x_60, x_60); +if (x_67 == 0) +{ +uint8_t x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_60); +lean_dec(x_59); +x_68 = 1; +x_69 = lean_box(x_68); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_69); +lean_ctor_set(x_9, 0, x_53); +x_70 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_70, 0, x_9); +x_18 = x_70; +x_19 = x_51; +goto block_24; +} +else +{ +size_t x_71; size_t x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; +x_71 = 0; +x_72 = lean_usize_of_nat(x_60); +lean_dec(x_60); +x_73 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_59, x_71, x_72, x_53); +lean_dec(x_59); +x_74 = 1; +x_75 = lean_box(x_74); +lean_ctor_set(x_26, 1, x_63); +lean_ctor_set(x_26, 0, x_75); +lean_ctor_set(x_9, 0, x_73); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_9); +x_18 = x_76; +x_19 = x_51; +goto block_24; +} +} +} +} +else +{ +uint8_t x_77; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_77 = !lean_is_exclusive(x_46); +if (x_77 == 0) +{ +return x_46; +} +else +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_46, 0); +x_79 = lean_ctor_get(x_46, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_46); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +} +else +{ +lean_object* x_81; +lean_dec(x_45); +lean_dec(x_34); +lean_dec(x_31); +x_81 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_81, 0, x_9); +x_18 = x_81; +x_19 = x_43; +goto block_24; +} +} +else +{ +uint8_t x_82; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_82 = !lean_is_exclusive(x_41); +if (x_82 == 0) +{ +return x_41; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_41, 0); +x_84 = lean_ctor_get(x_41, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_41); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +return x_85; +} +} +} +else +{ +lean_object* x_86; uint8_t x_87; +x_86 = lean_ctor_get(x_37, 1); +lean_inc(x_86); +lean_dec(x_37); +x_87 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_28, x_34); +if (x_87 == 0) +{ +lean_object* x_88; +lean_dec(x_31); +x_88 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_88, 0, x_9); +x_18 = x_88; +x_19 = x_86; +goto block_24; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_29); +x_89 = lean_box(0); +lean_inc(x_31); +x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_89); +x_91 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_89); +x_92 = 1; +x_93 = lean_box(x_92); +lean_ctor_set(x_26, 1, x_91); +lean_ctor_set(x_26, 0, x_93); +lean_ctor_set(x_9, 0, x_90); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_9); +x_18 = x_94; +x_19 = x_86; +goto block_24; +} +} +} +else +{ +uint8_t x_95; +lean_dec(x_34); +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_95 = !lean_is_exclusive(x_37); +if (x_95 == 0) +{ +return x_37; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_37, 0); +x_97 = lean_ctor_get(x_37, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_37); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +uint8_t x_99; +lean_dec(x_29); +x_99 = !lean_is_exclusive(x_36); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; +x_100 = lean_ctor_get(x_36, 0); +lean_dec(x_100); +x_101 = lean_box(0); +lean_inc(x_2); +x_102 = lean_array_mk(x_2); +x_103 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_104 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_101); +lean_ctor_set(x_104, 2, x_102); +x_105 = l_Lean_CollectFVars_main(x_34, x_104); +x_106 = lean_ctor_get(x_105, 2); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_array_get_size(x_106); +x_108 = lean_unsigned_to_nat(0u); +x_109 = lean_nat_dec_lt(x_108, x_107); +x_110 = lean_box(0); +x_111 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_110); +if (x_109 == 0) +{ +uint8_t x_112; lean_object* x_113; +lean_dec(x_107); +lean_dec(x_106); +x_112 = 1; +x_113 = lean_box(x_112); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_113); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +else +{ +uint8_t x_114; +x_114 = lean_nat_dec_le(x_107, x_107); +if (x_114 == 0) +{ +uint8_t x_115; lean_object* x_116; +lean_dec(x_107); +lean_dec(x_106); +x_115 = 1; +x_116 = lean_box(x_115); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_116); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +else +{ +size_t x_117; size_t x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; +x_117 = 0; +x_118 = lean_usize_of_nat(x_107); +lean_dec(x_107); +x_119 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_106, x_117, x_118, x_28); +lean_dec(x_106); +x_120 = 1; +x_121 = lean_box(x_120); +lean_ctor_set(x_26, 1, x_111); +lean_ctor_set(x_26, 0, x_121); +lean_ctor_set(x_9, 0, x_119); +lean_ctor_set(x_36, 0, x_9); +x_18 = x_36; +x_19 = x_35; +goto block_24; +} +} +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; +lean_dec(x_36); +x_122 = lean_box(0); +lean_inc(x_2); +x_123 = lean_array_mk(x_2); +x_124 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_125 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_122); +lean_ctor_set(x_125, 2, x_123); +x_126 = l_Lean_CollectFVars_main(x_34, x_125); +x_127 = lean_ctor_get(x_126, 2); +lean_inc(x_127); +lean_dec(x_126); +x_128 = lean_array_get_size(x_127); +x_129 = lean_unsigned_to_nat(0u); +x_130 = lean_nat_dec_lt(x_129, x_128); +x_131 = lean_box(0); +x_132 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_131); +if (x_130 == 0) +{ +uint8_t x_133; lean_object* x_134; lean_object* x_135; +lean_dec(x_128); +lean_dec(x_127); +x_133 = 1; +x_134 = lean_box(x_133); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_134); +x_135 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_135, 0, x_9); +x_18 = x_135; +x_19 = x_35; +goto block_24; +} +else +{ +uint8_t x_136; +x_136 = lean_nat_dec_le(x_128, x_128); +if (x_136 == 0) +{ +uint8_t x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_128); +lean_dec(x_127); +x_137 = 1; +x_138 = lean_box(x_137); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_138); +x_139 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_139, 0, x_9); +x_18 = x_139; +x_19 = x_35; +goto block_24; +} +else +{ +size_t x_140; size_t x_141; lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; +x_140 = 0; +x_141 = lean_usize_of_nat(x_128); +lean_dec(x_128); +x_142 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_127, x_140, x_141, x_28); +lean_dec(x_127); +x_143 = 1; +x_144 = lean_box(x_143); +lean_ctor_set(x_26, 1, x_132); +lean_ctor_set(x_26, 0, x_144); +lean_ctor_set(x_9, 0, x_142); +x_145 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_145, 0, x_9); +x_18 = x_145; +x_19 = x_35; +goto block_24; +} +} +} +} +} +else +{ +uint8_t x_146; +lean_dec(x_31); +lean_free_object(x_26); +lean_dec(x_30); +lean_dec(x_29); +lean_free_object(x_9); +lean_dec(x_28); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_146 = !lean_is_exclusive(x_33); +if (x_146 == 0) +{ +return x_33; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_33, 0); +x_148 = lean_ctor_get(x_33, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_33); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +return x_149; +} +} +} +else +{ +uint8_t x_150; +lean_dec(x_31); +lean_dec(x_17); +x_150 = !lean_is_exclusive(x_32); +if (x_150 == 0) +{ +lean_object* x_151; +x_151 = lean_ctor_get(x_32, 0); +lean_dec(x_151); +lean_ctor_set(x_32, 0, x_9); +x_18 = x_32; +x_19 = x_14; +goto block_24; +} +else +{ +lean_object* x_152; +lean_dec(x_32); +x_152 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_152, 0, x_9); +x_18 = x_152; +x_19 = x_14; +goto block_24; +} +} +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_153 = lean_ctor_get(x_9, 0); +x_154 = lean_ctor_get(x_26, 0); +x_155 = lean_ctor_get(x_26, 1); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_26); +x_156 = l_Lean_Expr_fvarId_x21(x_17); +x_157 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_155, x_156); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_158 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +lean_dec(x_158); +x_161 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_153, x_156); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_159); +x_162 = l_Lean_Meta_isProp(x_159, x_10, x_11, x_12, x_13, x_160); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_163; uint8_t x_164; +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_unbox(x_163); +lean_dec(x_163); +if (x_164 == 0) +{ +lean_object* x_165; lean_object* x_166; +x_165 = lean_ctor_get(x_162, 1); +lean_inc(x_165); +lean_dec(x_162); +lean_inc(x_10); +lean_inc(x_156); +x_166 = l_Lean_FVarId_getDecl(x_156, x_10, x_11, x_12, x_13, x_165); +if (lean_obj_tag(x_166) == 0) +{ +lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec(x_166); +x_169 = l_Lean_LocalDecl_binderInfo(x_167); +lean_dec(x_167); +x_170 = lean_box(x_169); +if (lean_obj_tag(x_170) == 3) +{ +lean_object* x_171; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_159); +lean_inc(x_153); +lean_inc(x_4); +x_171 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_153, x_159, x_10, x_11, x_12, x_13, x_168); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; uint8_t x_173; +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_unbox(x_172); +lean_dec(x_172); +if (x_173 == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; +lean_dec(x_159); +lean_dec(x_156); +x_174 = lean_ctor_get(x_171, 1); +lean_inc(x_174); +lean_dec(x_171); +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_154); +lean_ctor_set(x_175, 1, x_155); +lean_ctor_set(x_9, 1, x_175); +x_176 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_176, 0, x_9); +x_18 = x_176; +x_19 = x_174; +goto block_24; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; lean_object* x_189; +lean_dec(x_154); +x_177 = lean_ctor_get(x_171, 1); +lean_inc(x_177); +lean_dec(x_171); +x_178 = lean_box(0); +lean_inc(x_156); +x_179 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_178); +x_180 = lean_box(0); +lean_inc(x_2); +x_181 = lean_array_mk(x_2); +x_182 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_183 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_180); +lean_ctor_set(x_183, 2, x_181); +x_184 = l_Lean_CollectFVars_main(x_159, x_183); +x_185 = lean_ctor_get(x_184, 2); +lean_inc(x_185); +lean_dec(x_184); +x_186 = lean_array_get_size(x_185); +x_187 = lean_unsigned_to_nat(0u); +x_188 = lean_nat_dec_lt(x_187, x_186); +x_189 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_178); +if (x_188 == 0) +{ +uint8_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_186); +lean_dec(x_185); +x_190 = 1; +x_191 = lean_box(x_190); +x_192 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_192, 0, x_191); +lean_ctor_set(x_192, 1, x_189); +lean_ctor_set(x_9, 1, x_192); +lean_ctor_set(x_9, 0, x_179); +x_193 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_193, 0, x_9); +x_18 = x_193; +x_19 = x_177; +goto block_24; +} +else +{ +uint8_t x_194; +x_194 = lean_nat_dec_le(x_186, x_186); +if (x_194 == 0) +{ +uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_186); +lean_dec(x_185); +x_195 = 1; +x_196 = lean_box(x_195); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_189); +lean_ctor_set(x_9, 1, x_197); +lean_ctor_set(x_9, 0, x_179); +x_198 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_198, 0, x_9); +x_18 = x_198; +x_19 = x_177; +goto block_24; +} +else +{ +size_t x_199; size_t x_200; lean_object* x_201; uint8_t x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_199 = 0; +x_200 = lean_usize_of_nat(x_186); +lean_dec(x_186); +x_201 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_185, x_199, x_200, x_179); +lean_dec(x_185); +x_202 = 1; +x_203 = lean_box(x_202); +x_204 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_204, 0, x_203); +lean_ctor_set(x_204, 1, x_189); +lean_ctor_set(x_9, 1, x_204); +lean_ctor_set(x_9, 0, x_201); +x_205 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_205, 0, x_9); +x_18 = x_205; +x_19 = x_177; +goto block_24; +} +} +} +} +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_206 = lean_ctor_get(x_171, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_171, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_208 = x_171; +} else { + lean_dec_ref(x_171); + x_208 = lean_box(0); +} +if (lean_is_scalar(x_208)) { + x_209 = lean_alloc_ctor(1, 2, 0); +} else { + x_209 = x_208; +} +lean_ctor_set(x_209, 0, x_206); +lean_ctor_set(x_209, 1, x_207); +return x_209; +} +} +else +{ +lean_object* x_210; lean_object* x_211; +lean_dec(x_170); +lean_dec(x_159); +lean_dec(x_156); +x_210 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_210, 0, x_154); +lean_ctor_set(x_210, 1, x_155); +lean_ctor_set(x_9, 1, x_210); +x_211 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_211, 0, x_9); +x_18 = x_211; +x_19 = x_168; +goto block_24; +} +} +else +{ +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_212 = lean_ctor_get(x_166, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_166, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_214 = x_166; +} else { + lean_dec_ref(x_166); + x_214 = lean_box(0); +} +if (lean_is_scalar(x_214)) { + x_215 = lean_alloc_ctor(1, 2, 0); +} else { + x_215 = x_214; +} +lean_ctor_set(x_215, 0, x_212); +lean_ctor_set(x_215, 1, x_213); +return x_215; +} +} +else +{ +lean_object* x_216; uint8_t x_217; +x_216 = lean_ctor_get(x_162, 1); +lean_inc(x_216); +lean_dec(x_162); +x_217 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_153, x_159); +if (x_217 == 0) +{ +lean_object* x_218; lean_object* x_219; +lean_dec(x_156); +x_218 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_218, 0, x_154); +lean_ctor_set(x_218, 1, x_155); +lean_ctor_set(x_9, 1, x_218); +x_219 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_219, 0, x_9); +x_18 = x_219; +x_19 = x_216; +goto block_24; +} +else +{ +lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +lean_dec(x_154); +x_220 = lean_box(0); +lean_inc(x_156); +x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_220); +x_222 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_220); +x_223 = 1; +x_224 = lean_box(x_223); +x_225 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_225, 0, x_224); +lean_ctor_set(x_225, 1, x_222); +lean_ctor_set(x_9, 1, x_225); +lean_ctor_set(x_9, 0, x_221); +x_226 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_226, 0, x_9); +x_18 = x_226; +x_19 = x_216; +goto block_24; +} +} +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; +lean_dec(x_159); +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_227 = lean_ctor_get(x_162, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_162, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_162)) { + lean_ctor_release(x_162, 0); + lean_ctor_release(x_162, 1); + x_229 = x_162; +} else { + lean_dec_ref(x_162); + x_229 = lean_box(0); +} +if (lean_is_scalar(x_229)) { + x_230 = lean_alloc_ctor(1, 2, 0); +} else { + x_230 = x_229; +} +lean_ctor_set(x_230, 0, x_227); +lean_ctor_set(x_230, 1, x_228); +return x_230; +} +} +else +{ +lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; uint8_t x_240; lean_object* x_241; lean_object* x_242; +lean_dec(x_154); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + x_231 = x_161; +} else { + lean_dec_ref(x_161); + x_231 = lean_box(0); +} +x_232 = lean_box(0); +lean_inc(x_2); +x_233 = lean_array_mk(x_2); +x_234 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_235 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_235, 0, x_234); +lean_ctor_set(x_235, 1, x_232); +lean_ctor_set(x_235, 2, x_233); +x_236 = l_Lean_CollectFVars_main(x_159, x_235); +x_237 = lean_ctor_get(x_236, 2); +lean_inc(x_237); +lean_dec(x_236); +x_238 = lean_array_get_size(x_237); +x_239 = lean_unsigned_to_nat(0u); +x_240 = lean_nat_dec_lt(x_239, x_238); +x_241 = lean_box(0); +x_242 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_241); +if (x_240 == 0) +{ +uint8_t x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_dec(x_238); +lean_dec(x_237); +x_243 = 1; +x_244 = lean_box(x_243); +x_245 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_245, 0, x_244); +lean_ctor_set(x_245, 1, x_242); +lean_ctor_set(x_9, 1, x_245); +if (lean_is_scalar(x_231)) { + x_246 = lean_alloc_ctor(1, 1, 0); +} else { + x_246 = x_231; +} +lean_ctor_set(x_246, 0, x_9); +x_18 = x_246; +x_19 = x_160; +goto block_24; +} +else +{ +uint8_t x_247; +x_247 = lean_nat_dec_le(x_238, x_238); +if (x_247 == 0) +{ +uint8_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; +lean_dec(x_238); +lean_dec(x_237); +x_248 = 1; +x_249 = lean_box(x_248); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_242); +lean_ctor_set(x_9, 1, x_250); +if (lean_is_scalar(x_231)) { + x_251 = lean_alloc_ctor(1, 1, 0); +} else { + x_251 = x_231; +} +lean_ctor_set(x_251, 0, x_9); +x_18 = x_251; +x_19 = x_160; +goto block_24; +} +else +{ +size_t x_252; size_t x_253; lean_object* x_254; uint8_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; +x_252 = 0; +x_253 = lean_usize_of_nat(x_238); +lean_dec(x_238); +x_254 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_237, x_252, x_253, x_153); +lean_dec(x_237); +x_255 = 1; +x_256 = lean_box(x_255); +x_257 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_257, 0, x_256); +lean_ctor_set(x_257, 1, x_242); +lean_ctor_set(x_9, 1, x_257); +lean_ctor_set(x_9, 0, x_254); +if (lean_is_scalar(x_231)) { + x_258 = lean_alloc_ctor(1, 1, 0); +} else { + x_258 = x_231; +} +lean_ctor_set(x_258, 0, x_9); +x_18 = x_258; +x_19 = x_160; +goto block_24; +} +} +} +} +else +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; +lean_dec(x_156); +lean_dec(x_155); +lean_dec(x_154); +lean_free_object(x_9); +lean_dec(x_153); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_259 = lean_ctor_get(x_158, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_158, 1); +lean_inc(x_260); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_261 = x_158; +} else { + lean_dec_ref(x_158); + x_261 = lean_box(0); +} +if (lean_is_scalar(x_261)) { + x_262 = lean_alloc_ctor(1, 2, 0); +} else { + x_262 = x_261; +} +lean_ctor_set(x_262, 0, x_259); +lean_ctor_set(x_262, 1, x_260); +return x_262; +} +} +else +{ +lean_object* x_263; lean_object* x_264; lean_object* x_265; +lean_dec(x_156); +lean_dec(x_17); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + x_263 = x_157; +} else { + lean_dec_ref(x_157); + x_263 = lean_box(0); +} +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_154); +lean_ctor_set(x_264, 1, x_155); +lean_ctor_set(x_9, 1, x_264); +if (lean_is_scalar(x_263)) { + x_265 = lean_alloc_ctor(1, 1, 0); +} else { + x_265 = x_263; +} +lean_ctor_set(x_265, 0, x_9); +x_18 = x_265; +x_19 = x_14; +goto block_24; +} +} +} +else +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; +x_266 = lean_ctor_get(x_9, 1); +x_267 = lean_ctor_get(x_9, 0); +lean_inc(x_266); +lean_inc(x_267); +lean_dec(x_9); +x_268 = lean_ctor_get(x_266, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_266, 1); +lean_inc(x_269); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_270 = x_266; +} else { + lean_dec_ref(x_266); + x_270 = lean_box(0); +} +x_271 = l_Lean_Expr_fvarId_x21(x_17); +x_272 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_269, x_271); +if (lean_obj_tag(x_272) == 0) +{ +lean_object* x_273; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_273 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_273) == 0) +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; +x_274 = lean_ctor_get(x_273, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_273, 1); +lean_inc(x_275); +lean_dec(x_273); +x_276 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_267, x_271); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_274); +x_277 = l_Lean_Meta_isProp(x_274, x_10, x_11, x_12, x_13, x_275); +if (lean_obj_tag(x_277) == 0) +{ +lean_object* x_278; uint8_t x_279; +x_278 = lean_ctor_get(x_277, 0); +lean_inc(x_278); +x_279 = lean_unbox(x_278); +lean_dec(x_278); +if (x_279 == 0) +{ +lean_object* x_280; lean_object* x_281; +x_280 = lean_ctor_get(x_277, 1); +lean_inc(x_280); +lean_dec(x_277); +lean_inc(x_10); +lean_inc(x_271); +x_281 = l_Lean_FVarId_getDecl(x_271, x_10, x_11, x_12, x_13, x_280); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; uint8_t x_284; lean_object* x_285; +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_284 = l_Lean_LocalDecl_binderInfo(x_282); +lean_dec(x_282); +x_285 = lean_box(x_284); +if (lean_obj_tag(x_285) == 3) +{ +lean_object* x_286; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_274); +lean_inc(x_267); +lean_inc(x_4); +x_286 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_267, x_274, x_10, x_11, x_12, x_13, x_283); +if (lean_obj_tag(x_286) == 0) +{ +lean_object* x_287; uint8_t x_288; +x_287 = lean_ctor_get(x_286, 0); +lean_inc(x_287); +x_288 = lean_unbox(x_287); +lean_dec(x_287); +if (x_288 == 0) +{ +lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; +lean_dec(x_274); +lean_dec(x_271); +x_289 = lean_ctor_get(x_286, 1); +lean_inc(x_289); +lean_dec(x_286); +if (lean_is_scalar(x_270)) { + x_290 = lean_alloc_ctor(0, 2, 0); +} else { + x_290 = x_270; +} +lean_ctor_set(x_290, 0, x_268); +lean_ctor_set(x_290, 1, x_269); +x_291 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_291, 0, x_267); +lean_ctor_set(x_291, 1, x_290); +x_292 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_292, 0, x_291); +x_18 = x_292; +x_19 = x_289; +goto block_24; +} +else +{ +lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; lean_object* x_305; +lean_dec(x_268); +x_293 = lean_ctor_get(x_286, 1); +lean_inc(x_293); +lean_dec(x_286); +x_294 = lean_box(0); +lean_inc(x_271); +x_295 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_294); +x_296 = lean_box(0); +lean_inc(x_2); +x_297 = lean_array_mk(x_2); +x_298 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_299 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_299, 0, x_298); +lean_ctor_set(x_299, 1, x_296); +lean_ctor_set(x_299, 2, x_297); +x_300 = l_Lean_CollectFVars_main(x_274, x_299); +x_301 = lean_ctor_get(x_300, 2); +lean_inc(x_301); +lean_dec(x_300); +x_302 = lean_array_get_size(x_301); +x_303 = lean_unsigned_to_nat(0u); +x_304 = lean_nat_dec_lt(x_303, x_302); +x_305 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_294); +if (x_304 == 0) +{ +uint8_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +lean_dec(x_302); +lean_dec(x_301); +x_306 = 1; +x_307 = lean_box(x_306); +if (lean_is_scalar(x_270)) { + x_308 = lean_alloc_ctor(0, 2, 0); +} else { + x_308 = x_270; +} +lean_ctor_set(x_308, 0, x_307); +lean_ctor_set(x_308, 1, x_305); +x_309 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_309, 0, x_295); +lean_ctor_set(x_309, 1, x_308); +x_310 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_310, 0, x_309); +x_18 = x_310; +x_19 = x_293; +goto block_24; +} +else +{ +uint8_t x_311; +x_311 = lean_nat_dec_le(x_302, x_302); +if (x_311 == 0) +{ +uint8_t x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; +lean_dec(x_302); +lean_dec(x_301); +x_312 = 1; +x_313 = lean_box(x_312); +if (lean_is_scalar(x_270)) { + x_314 = lean_alloc_ctor(0, 2, 0); +} else { + x_314 = x_270; +} +lean_ctor_set(x_314, 0, x_313); +lean_ctor_set(x_314, 1, x_305); +x_315 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_315, 0, x_295); +lean_ctor_set(x_315, 1, x_314); +x_316 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_316, 0, x_315); +x_18 = x_316; +x_19 = x_293; +goto block_24; +} +else +{ +size_t x_317; size_t x_318; lean_object* x_319; uint8_t x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +x_317 = 0; +x_318 = lean_usize_of_nat(x_302); +lean_dec(x_302); +x_319 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_301, x_317, x_318, x_295); +lean_dec(x_301); +x_320 = 1; +x_321 = lean_box(x_320); +if (lean_is_scalar(x_270)) { + x_322 = lean_alloc_ctor(0, 2, 0); +} else { + x_322 = x_270; +} +lean_ctor_set(x_322, 0, x_321); +lean_ctor_set(x_322, 1, x_305); +x_323 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_323, 0, x_319); +lean_ctor_set(x_323, 1, x_322); +x_324 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_324, 0, x_323); +x_18 = x_324; +x_19 = x_293; +goto block_24; +} +} +} +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_325 = lean_ctor_get(x_286, 0); +lean_inc(x_325); +x_326 = lean_ctor_get(x_286, 1); +lean_inc(x_326); +if (lean_is_exclusive(x_286)) { + lean_ctor_release(x_286, 0); + lean_ctor_release(x_286, 1); + x_327 = x_286; +} else { + lean_dec_ref(x_286); + x_327 = lean_box(0); +} +if (lean_is_scalar(x_327)) { + x_328 = lean_alloc_ctor(1, 2, 0); +} else { + x_328 = x_327; +} +lean_ctor_set(x_328, 0, x_325); +lean_ctor_set(x_328, 1, x_326); +return x_328; +} +} +else +{ +lean_object* x_329; lean_object* x_330; lean_object* x_331; +lean_dec(x_285); +lean_dec(x_274); +lean_dec(x_271); +if (lean_is_scalar(x_270)) { + x_329 = lean_alloc_ctor(0, 2, 0); +} else { + x_329 = x_270; +} +lean_ctor_set(x_329, 0, x_268); +lean_ctor_set(x_329, 1, x_269); +x_330 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_330, 0, x_267); +lean_ctor_set(x_330, 1, x_329); +x_331 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_331, 0, x_330); +x_18 = x_331; +x_19 = x_283; +goto block_24; +} +} +else +{ +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_332 = lean_ctor_get(x_281, 0); +lean_inc(x_332); +x_333 = lean_ctor_get(x_281, 1); +lean_inc(x_333); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_334 = x_281; +} else { + lean_dec_ref(x_281); + x_334 = lean_box(0); +} +if (lean_is_scalar(x_334)) { + x_335 = lean_alloc_ctor(1, 2, 0); +} else { + x_335 = x_334; +} +lean_ctor_set(x_335, 0, x_332); +lean_ctor_set(x_335, 1, x_333); +return x_335; +} +} +else +{ +lean_object* x_336; uint8_t x_337; +x_336 = lean_ctor_get(x_277, 1); +lean_inc(x_336); +lean_dec(x_277); +x_337 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_267, x_274); +if (x_337 == 0) +{ +lean_object* x_338; lean_object* x_339; lean_object* x_340; +lean_dec(x_271); +if (lean_is_scalar(x_270)) { + x_338 = lean_alloc_ctor(0, 2, 0); +} else { + x_338 = x_270; +} +lean_ctor_set(x_338, 0, x_268); +lean_ctor_set(x_338, 1, x_269); +x_339 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_339, 0, x_267); +lean_ctor_set(x_339, 1, x_338); +x_340 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_340, 0, x_339); +x_18 = x_340; +x_19 = x_336; +goto block_24; +} +else +{ +lean_object* x_341; lean_object* x_342; lean_object* x_343; uint8_t x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; +lean_dec(x_268); +x_341 = lean_box(0); +lean_inc(x_271); +x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_341); +x_343 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_341); +x_344 = 1; +x_345 = lean_box(x_344); +if (lean_is_scalar(x_270)) { + x_346 = lean_alloc_ctor(0, 2, 0); +} else { + x_346 = x_270; +} +lean_ctor_set(x_346, 0, x_345); +lean_ctor_set(x_346, 1, x_343); +x_347 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_347, 0, x_342); +lean_ctor_set(x_347, 1, x_346); +x_348 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_348, 0, x_347); +x_18 = x_348; +x_19 = x_336; +goto block_24; +} +} +} +else +{ +lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; +lean_dec(x_274); +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_349 = lean_ctor_get(x_277, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_277, 1); +lean_inc(x_350); +if (lean_is_exclusive(x_277)) { + lean_ctor_release(x_277, 0); + lean_ctor_release(x_277, 1); + x_351 = x_277; +} else { + lean_dec_ref(x_277); + x_351 = lean_box(0); +} +if (lean_is_scalar(x_351)) { + x_352 = lean_alloc_ctor(1, 2, 0); +} else { + x_352 = x_351; +} +lean_ctor_set(x_352, 0, x_349); +lean_ctor_set(x_352, 1, x_350); +return x_352; +} +} +else +{ +lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; uint8_t x_362; lean_object* x_363; lean_object* x_364; +lean_dec(x_268); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + x_353 = x_276; +} else { + lean_dec_ref(x_276); + x_353 = lean_box(0); +} +x_354 = lean_box(0); +lean_inc(x_2); +x_355 = lean_array_mk(x_2); +x_356 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_357 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_357, 0, x_356); +lean_ctor_set(x_357, 1, x_354); +lean_ctor_set(x_357, 2, x_355); +x_358 = l_Lean_CollectFVars_main(x_274, x_357); +x_359 = lean_ctor_get(x_358, 2); +lean_inc(x_359); +lean_dec(x_358); +x_360 = lean_array_get_size(x_359); +x_361 = lean_unsigned_to_nat(0u); +x_362 = lean_nat_dec_lt(x_361, x_360); +x_363 = lean_box(0); +x_364 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_363); +if (x_362 == 0) +{ +uint8_t x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +lean_dec(x_360); +lean_dec(x_359); +x_365 = 1; +x_366 = lean_box(x_365); +if (lean_is_scalar(x_270)) { + x_367 = lean_alloc_ctor(0, 2, 0); +} else { + x_367 = x_270; +} +lean_ctor_set(x_367, 0, x_366); +lean_ctor_set(x_367, 1, x_364); +x_368 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_368, 0, x_267); +lean_ctor_set(x_368, 1, x_367); +if (lean_is_scalar(x_353)) { + x_369 = lean_alloc_ctor(1, 1, 0); +} else { + x_369 = x_353; +} +lean_ctor_set(x_369, 0, x_368); +x_18 = x_369; +x_19 = x_275; +goto block_24; +} +else +{ +uint8_t x_370; +x_370 = lean_nat_dec_le(x_360, x_360); +if (x_370 == 0) +{ +uint8_t x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +lean_dec(x_360); +lean_dec(x_359); +x_371 = 1; +x_372 = lean_box(x_371); +if (lean_is_scalar(x_270)) { + x_373 = lean_alloc_ctor(0, 2, 0); +} else { + x_373 = x_270; +} +lean_ctor_set(x_373, 0, x_372); +lean_ctor_set(x_373, 1, x_364); +x_374 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_374, 0, x_267); +lean_ctor_set(x_374, 1, x_373); +if (lean_is_scalar(x_353)) { + x_375 = lean_alloc_ctor(1, 1, 0); +} else { + x_375 = x_353; +} +lean_ctor_set(x_375, 0, x_374); +x_18 = x_375; +x_19 = x_275; +goto block_24; +} +else +{ +size_t x_376; size_t x_377; lean_object* x_378; uint8_t x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; +x_376 = 0; +x_377 = lean_usize_of_nat(x_360); +lean_dec(x_360); +x_378 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_359, x_376, x_377, x_267); +lean_dec(x_359); +x_379 = 1; +x_380 = lean_box(x_379); +if (lean_is_scalar(x_270)) { + x_381 = lean_alloc_ctor(0, 2, 0); +} else { + x_381 = x_270; +} +lean_ctor_set(x_381, 0, x_380); +lean_ctor_set(x_381, 1, x_364); +x_382 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_382, 0, x_378); +lean_ctor_set(x_382, 1, x_381); +if (lean_is_scalar(x_353)) { + x_383 = lean_alloc_ctor(1, 1, 0); +} else { + x_383 = x_353; +} +lean_ctor_set(x_383, 0, x_382); +x_18 = x_383; +x_19 = x_275; +goto block_24; +} +} +} +} +else +{ +lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +lean_dec(x_271); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_2); +x_384 = lean_ctor_get(x_273, 0); +lean_inc(x_384); +x_385 = lean_ctor_get(x_273, 1); +lean_inc(x_385); +if (lean_is_exclusive(x_273)) { + lean_ctor_release(x_273, 0); + lean_ctor_release(x_273, 1); + x_386 = x_273; +} else { + lean_dec_ref(x_273); + x_386 = lean_box(0); +} +if (lean_is_scalar(x_386)) { + x_387 = lean_alloc_ctor(1, 2, 0); +} else { + x_387 = x_386; +} +lean_ctor_set(x_387, 0, x_384); +lean_ctor_set(x_387, 1, x_385); +return x_387; +} +} +else +{ +lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; +lean_dec(x_271); +lean_dec(x_17); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + x_388 = x_272; +} else { + lean_dec_ref(x_272); + x_388 = lean_box(0); +} +if (lean_is_scalar(x_270)) { + x_389 = lean_alloc_ctor(0, 2, 0); +} else { + x_389 = x_270; +} +lean_ctor_set(x_389, 0, x_268); +lean_ctor_set(x_389, 1, x_269); +x_390 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_390, 0, x_267); +lean_ctor_set(x_390, 1, x_389); +if (lean_is_scalar(x_388)) { + x_391 = lean_alloc_ctor(1, 1, 0); +} else { + x_391 = x_388; +} +lean_ctor_set(x_391, 0, x_390); +x_18 = x_391; +x_19 = x_14; +goto block_24; +} +} +block_24: +{ +lean_object* x_20; size_t x_21; size_t x_22; +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = 1; +x_22 = lean_usize_add(x_8, x_21); +x_8 = x_22; +x_9 = x_20; +x_14 = x_19; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_7, x_6); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_24; +x_16 = lean_array_uget(x_5, x_7); +x_24 = !lean_is_exclusive(x_8); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_8, 1); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_8, 0); +x_28 = lean_ctor_get(x_25, 0); +x_29 = lean_ctor_get(x_25, 1); +x_30 = l_Lean_Expr_fvarId_x21(x_16); +x_31 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_29, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_32 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_27, x_30); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_33); +x_36 = l_Lean_Meta_isProp(x_33, x_9, x_10, x_11, x_12, x_34); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; uint8_t x_38; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_unbox(x_37); +lean_dec(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +lean_inc(x_9); +lean_inc(x_30); +x_40 = l_Lean_FVarId_getDecl(x_30, x_9, x_10, x_11, x_12, x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = l_Lean_LocalDecl_binderInfo(x_41); +lean_dec(x_41); +x_44 = lean_box(x_43); +if (lean_obj_tag(x_44) == 3) +{ +lean_object* x_45; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_33); +lean_inc(x_27); +lean_inc(x_3); +x_45 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_27, x_33, x_9, x_10, x_11, x_12, x_42); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_33); +lean_dec(x_30); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +x_49 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_49, 0, x_8); +x_17 = x_49; +x_18 = x_48; +goto block_23; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; +lean_dec(x_28); +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_dec(x_45); +x_51 = lean_box(0); +lean_inc(x_30); +x_52 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_51); +x_53 = lean_box(0); +lean_inc(x_2); +x_54 = lean_array_mk(x_2); +x_55 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +lean_ctor_set(x_56, 2, x_54); +x_57 = l_Lean_CollectFVars_main(x_33, x_56); +x_58 = lean_ctor_get(x_57, 2); +lean_inc(x_58); +lean_dec(x_57); +x_59 = lean_array_get_size(x_58); +x_60 = lean_unsigned_to_nat(0u); +x_61 = lean_nat_dec_lt(x_60, x_59); +x_62 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_51); +if (x_61 == 0) +{ +uint8_t x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_59); +lean_dec(x_58); +x_63 = 1; +x_64 = lean_box(x_63); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_64); +lean_ctor_set(x_8, 0, x_52); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_8); +x_17 = x_65; +x_18 = x_50; +goto block_23; +} +else +{ +uint8_t x_66; +x_66 = lean_nat_dec_le(x_59, x_59); +if (x_66 == 0) +{ +uint8_t x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_59); +lean_dec(x_58); +x_67 = 1; +x_68 = lean_box(x_67); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_68); +lean_ctor_set(x_8, 0, x_52); +x_69 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_69, 0, x_8); +x_17 = x_69; +x_18 = x_50; +goto block_23; +} +else +{ +size_t x_70; size_t x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; +x_70 = 0; +x_71 = lean_usize_of_nat(x_59); +lean_dec(x_59); +x_72 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_58, x_70, x_71, x_52); +lean_dec(x_58); +x_73 = 1; +x_74 = lean_box(x_73); +lean_ctor_set(x_25, 1, x_62); +lean_ctor_set(x_25, 0, x_74); +lean_ctor_set(x_8, 0, x_72); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_8); +x_17 = x_75; +x_18 = x_50; +goto block_23; +} +} +} +} +else +{ +uint8_t x_76; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_76 = !lean_is_exclusive(x_45); +if (x_76 == 0) +{ +return x_45; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_45, 0); +x_78 = lean_ctor_get(x_45, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_45); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} +} +else +{ +lean_object* x_80; +lean_dec(x_44); +lean_dec(x_33); +lean_dec(x_30); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_8); +x_17 = x_80; +x_18 = x_42; +goto block_23; +} +} +else +{ +uint8_t x_81; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_81 = !lean_is_exclusive(x_40); +if (x_81 == 0) +{ +return x_40; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_40, 0); +x_83 = lean_ctor_get(x_40, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_40); +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +return x_84; +} +} +} +else +{ +lean_object* x_85; uint8_t x_86; +x_85 = lean_ctor_get(x_36, 1); +lean_inc(x_85); +lean_dec(x_36); +x_86 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_27, x_33); +if (x_86 == 0) +{ +lean_object* x_87; +lean_dec(x_30); +x_87 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_87, 0, x_8); +x_17 = x_87; +x_18 = x_85; +goto block_23; +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_28); +x_88 = lean_box(0); +lean_inc(x_30); +x_89 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_88); +x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_88); +x_91 = 1; +x_92 = lean_box(x_91); +lean_ctor_set(x_25, 1, x_90); +lean_ctor_set(x_25, 0, x_92); +lean_ctor_set(x_8, 0, x_89); +x_93 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_93, 0, x_8); +x_17 = x_93; +x_18 = x_85; +goto block_23; +} +} +} +else +{ +uint8_t x_94; +lean_dec(x_33); +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_94 = !lean_is_exclusive(x_36); +if (x_94 == 0) +{ +return x_36; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_36, 0); +x_96 = lean_ctor_get(x_36, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_36); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; +} +} +} +else +{ +uint8_t x_98; +lean_dec(x_28); +x_98 = !lean_is_exclusive(x_35); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; +x_99 = lean_ctor_get(x_35, 0); +lean_dec(x_99); +x_100 = lean_box(0); +lean_inc(x_2); +x_101 = lean_array_mk(x_2); +x_102 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_103 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_100); +lean_ctor_set(x_103, 2, x_101); +x_104 = l_Lean_CollectFVars_main(x_33, x_103); +x_105 = lean_ctor_get(x_104, 2); +lean_inc(x_105); +lean_dec(x_104); +x_106 = lean_array_get_size(x_105); +x_107 = lean_unsigned_to_nat(0u); +x_108 = lean_nat_dec_lt(x_107, x_106); +x_109 = lean_box(0); +x_110 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_109); +if (x_108 == 0) +{ +uint8_t x_111; lean_object* x_112; +lean_dec(x_106); +lean_dec(x_105); +x_111 = 1; +x_112 = lean_box(x_111); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_112); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +else +{ +uint8_t x_113; +x_113 = lean_nat_dec_le(x_106, x_106); +if (x_113 == 0) +{ +uint8_t x_114; lean_object* x_115; +lean_dec(x_106); +lean_dec(x_105); +x_114 = 1; +x_115 = lean_box(x_114); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_115); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +else +{ +size_t x_116; size_t x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; +x_116 = 0; +x_117 = lean_usize_of_nat(x_106); +lean_dec(x_106); +x_118 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_105, x_116, x_117, x_27); +lean_dec(x_105); +x_119 = 1; +x_120 = lean_box(x_119); +lean_ctor_set(x_25, 1, x_110); +lean_ctor_set(x_25, 0, x_120); +lean_ctor_set(x_8, 0, x_118); +lean_ctor_set(x_35, 0, x_8); +x_17 = x_35; +x_18 = x_34; +goto block_23; +} +} +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; +lean_dec(x_35); +x_121 = lean_box(0); +lean_inc(x_2); +x_122 = lean_array_mk(x_2); +x_123 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_124 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_121); +lean_ctor_set(x_124, 2, x_122); +x_125 = l_Lean_CollectFVars_main(x_33, x_124); +x_126 = lean_ctor_get(x_125, 2); +lean_inc(x_126); +lean_dec(x_125); +x_127 = lean_array_get_size(x_126); +x_128 = lean_unsigned_to_nat(0u); +x_129 = lean_nat_dec_lt(x_128, x_127); +x_130 = lean_box(0); +x_131 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_130); +if (x_129 == 0) +{ +uint8_t x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_127); +lean_dec(x_126); +x_132 = 1; +x_133 = lean_box(x_132); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_133); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_8); +x_17 = x_134; +x_18 = x_34; +goto block_23; +} +else +{ +uint8_t x_135; +x_135 = lean_nat_dec_le(x_127, x_127); +if (x_135 == 0) +{ +uint8_t x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_127); +lean_dec(x_126); +x_136 = 1; +x_137 = lean_box(x_136); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_137); +x_138 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_138, 0, x_8); +x_17 = x_138; +x_18 = x_34; +goto block_23; +} +else +{ +size_t x_139; size_t x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; +x_139 = 0; +x_140 = lean_usize_of_nat(x_127); +lean_dec(x_127); +x_141 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_126, x_139, x_140, x_27); +lean_dec(x_126); +x_142 = 1; +x_143 = lean_box(x_142); +lean_ctor_set(x_25, 1, x_131); +lean_ctor_set(x_25, 0, x_143); +lean_ctor_set(x_8, 0, x_141); +x_144 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_144, 0, x_8); +x_17 = x_144; +x_18 = x_34; +goto block_23; +} +} +} +} +} +else +{ +uint8_t x_145; +lean_dec(x_30); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_28); +lean_free_object(x_8); +lean_dec(x_27); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_145 = !lean_is_exclusive(x_32); +if (x_145 == 0) +{ +return x_32; +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_32, 0); +x_147 = lean_ctor_get(x_32, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_32); +x_148 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +return x_148; +} +} +} +else +{ +uint8_t x_149; +lean_dec(x_30); +lean_dec(x_16); +x_149 = !lean_is_exclusive(x_31); +if (x_149 == 0) +{ +lean_object* x_150; +x_150 = lean_ctor_get(x_31, 0); +lean_dec(x_150); +lean_ctor_set(x_31, 0, x_8); +x_17 = x_31; +x_18 = x_13; +goto block_23; +} +else +{ +lean_object* x_151; +lean_dec(x_31); +x_151 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_151, 0, x_8); +x_17 = x_151; +x_18 = x_13; +goto block_23; +} +} +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_152 = lean_ctor_get(x_8, 0); +x_153 = lean_ctor_get(x_25, 0); +x_154 = lean_ctor_get(x_25, 1); +lean_inc(x_154); +lean_inc(x_153); +lean_dec(x_25); +x_155 = l_Lean_Expr_fvarId_x21(x_16); +x_156 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_154, x_155); +if (lean_obj_tag(x_156) == 0) +{ +lean_object* x_157; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_157 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_157, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_157, 1); +lean_inc(x_159); +lean_dec(x_157); +x_160 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_152, x_155); +if (lean_obj_tag(x_160) == 0) +{ +lean_object* x_161; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_158); +x_161 = l_Lean_Meta_isProp(x_158, x_9, x_10, x_11, x_12, x_159); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; uint8_t x_163; +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = lean_unbox(x_162); +lean_dec(x_162); +if (x_163 == 0) +{ +lean_object* x_164; lean_object* x_165; +x_164 = lean_ctor_get(x_161, 1); +lean_inc(x_164); +lean_dec(x_161); +lean_inc(x_9); +lean_inc(x_155); +x_165 = l_Lean_FVarId_getDecl(x_155, x_9, x_10, x_11, x_12, x_164); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; +x_166 = lean_ctor_get(x_165, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_165, 1); +lean_inc(x_167); +lean_dec(x_165); +x_168 = l_Lean_LocalDecl_binderInfo(x_166); +lean_dec(x_166); +x_169 = lean_box(x_168); +if (lean_obj_tag(x_169) == 3) +{ +lean_object* x_170; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_158); +lean_inc(x_152); +lean_inc(x_3); +x_170 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_152, x_158, x_9, x_10, x_11, x_12, x_167); +if (lean_obj_tag(x_170) == 0) +{ +lean_object* x_171; uint8_t x_172; +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +x_172 = lean_unbox(x_171); +lean_dec(x_171); +if (x_172 == 0) +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_158); +lean_dec(x_155); +x_173 = lean_ctor_get(x_170, 1); +lean_inc(x_173); +lean_dec(x_170); +x_174 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_174, 0, x_153); +lean_ctor_set(x_174, 1, x_154); +lean_ctor_set(x_8, 1, x_174); +x_175 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_175, 0, x_8); +x_17 = x_175; +x_18 = x_173; +goto block_23; +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; +lean_dec(x_153); +x_176 = lean_ctor_get(x_170, 1); +lean_inc(x_176); +lean_dec(x_170); +x_177 = lean_box(0); +lean_inc(x_155); +x_178 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_177); +x_179 = lean_box(0); +lean_inc(x_2); +x_180 = lean_array_mk(x_2); +x_181 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_182 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_182, 0, x_181); +lean_ctor_set(x_182, 1, x_179); +lean_ctor_set(x_182, 2, x_180); +x_183 = l_Lean_CollectFVars_main(x_158, x_182); +x_184 = lean_ctor_get(x_183, 2); +lean_inc(x_184); +lean_dec(x_183); +x_185 = lean_array_get_size(x_184); +x_186 = lean_unsigned_to_nat(0u); +x_187 = lean_nat_dec_lt(x_186, x_185); +x_188 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_177); +if (x_187 == 0) +{ +uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +lean_dec(x_185); +lean_dec(x_184); +x_189 = 1; +x_190 = lean_box(x_189); +x_191 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_188); +lean_ctor_set(x_8, 1, x_191); +lean_ctor_set(x_8, 0, x_178); +x_192 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_192, 0, x_8); +x_17 = x_192; +x_18 = x_176; +goto block_23; +} +else +{ +uint8_t x_193; +x_193 = lean_nat_dec_le(x_185, x_185); +if (x_193 == 0) +{ +uint8_t x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +lean_dec(x_185); +lean_dec(x_184); +x_194 = 1; +x_195 = lean_box(x_194); +x_196 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_196, 0, x_195); +lean_ctor_set(x_196, 1, x_188); +lean_ctor_set(x_8, 1, x_196); +lean_ctor_set(x_8, 0, x_178); +x_197 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_197, 0, x_8); +x_17 = x_197; +x_18 = x_176; +goto block_23; +} +else +{ +size_t x_198; size_t x_199; lean_object* x_200; uint8_t x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_198 = 0; +x_199 = lean_usize_of_nat(x_185); +lean_dec(x_185); +x_200 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_184, x_198, x_199, x_178); +lean_dec(x_184); +x_201 = 1; +x_202 = lean_box(x_201); +x_203 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_188); +lean_ctor_set(x_8, 1, x_203); +lean_ctor_set(x_8, 0, x_200); +x_204 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_204, 0, x_8); +x_17 = x_204; +x_18 = x_176; +goto block_23; +} +} +} +} +else +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_205 = lean_ctor_get(x_170, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_170, 1); +lean_inc(x_206); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_207 = x_170; +} else { + lean_dec_ref(x_170); + x_207 = lean_box(0); +} +if (lean_is_scalar(x_207)) { + x_208 = lean_alloc_ctor(1, 2, 0); +} else { + x_208 = x_207; +} +lean_ctor_set(x_208, 0, x_205); +lean_ctor_set(x_208, 1, x_206); +return x_208; +} +} +else +{ +lean_object* x_209; lean_object* x_210; +lean_dec(x_169); +lean_dec(x_158); +lean_dec(x_155); +x_209 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_209, 0, x_153); +lean_ctor_set(x_209, 1, x_154); +lean_ctor_set(x_8, 1, x_209); +x_210 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_210, 0, x_8); +x_17 = x_210; +x_18 = x_167; +goto block_23; +} +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_211 = lean_ctor_get(x_165, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_165, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_165)) { + lean_ctor_release(x_165, 0); + lean_ctor_release(x_165, 1); + x_213 = x_165; +} else { + lean_dec_ref(x_165); + x_213 = lean_box(0); +} +if (lean_is_scalar(x_213)) { + x_214 = lean_alloc_ctor(1, 2, 0); +} else { + x_214 = x_213; +} +lean_ctor_set(x_214, 0, x_211); +lean_ctor_set(x_214, 1, x_212); +return x_214; +} +} +else +{ +lean_object* x_215; uint8_t x_216; +x_215 = lean_ctor_get(x_161, 1); +lean_inc(x_215); +lean_dec(x_161); +x_216 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_152, x_158); +if (x_216 == 0) +{ +lean_object* x_217; lean_object* x_218; +lean_dec(x_155); +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_153); +lean_ctor_set(x_217, 1, x_154); +lean_ctor_set(x_8, 1, x_217); +x_218 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_218, 0, x_8); +x_17 = x_218; +x_18 = x_215; +goto block_23; +} +else +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; uint8_t x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; +lean_dec(x_153); +x_219 = lean_box(0); +lean_inc(x_155); +x_220 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_219); +x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_219); +x_222 = 1; +x_223 = lean_box(x_222); +x_224 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_224, 0, x_223); +lean_ctor_set(x_224, 1, x_221); +lean_ctor_set(x_8, 1, x_224); +lean_ctor_set(x_8, 0, x_220); +x_225 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_225, 0, x_8); +x_17 = x_225; +x_18 = x_215; +goto block_23; +} +} +} +else +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_158); +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_226 = lean_ctor_get(x_161, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_161, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_228 = x_161; +} else { + lean_dec_ref(x_161); + x_228 = lean_box(0); +} +if (lean_is_scalar(x_228)) { + x_229 = lean_alloc_ctor(1, 2, 0); +} else { + x_229 = x_228; +} +lean_ctor_set(x_229, 0, x_226); +lean_ctor_set(x_229, 1, x_227); +return x_229; +} +} +else +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; lean_object* x_240; lean_object* x_241; +lean_dec(x_153); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + x_230 = x_160; +} else { + lean_dec_ref(x_160); + x_230 = lean_box(0); +} +x_231 = lean_box(0); +lean_inc(x_2); +x_232 = lean_array_mk(x_2); +x_233 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_234 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_234, 0, x_233); +lean_ctor_set(x_234, 1, x_231); +lean_ctor_set(x_234, 2, x_232); +x_235 = l_Lean_CollectFVars_main(x_158, x_234); +x_236 = lean_ctor_get(x_235, 2); +lean_inc(x_236); +lean_dec(x_235); +x_237 = lean_array_get_size(x_236); +x_238 = lean_unsigned_to_nat(0u); +x_239 = lean_nat_dec_lt(x_238, x_237); +x_240 = lean_box(0); +x_241 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_240); +if (x_239 == 0) +{ +uint8_t x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +lean_dec(x_237); +lean_dec(x_236); +x_242 = 1; +x_243 = lean_box(x_242); +x_244 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_241); +lean_ctor_set(x_8, 1, x_244); +if (lean_is_scalar(x_230)) { + x_245 = lean_alloc_ctor(1, 1, 0); +} else { + x_245 = x_230; +} +lean_ctor_set(x_245, 0, x_8); +x_17 = x_245; +x_18 = x_159; +goto block_23; +} +else +{ +uint8_t x_246; +x_246 = lean_nat_dec_le(x_237, x_237); +if (x_246 == 0) +{ +uint8_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +lean_dec(x_237); +lean_dec(x_236); +x_247 = 1; +x_248 = lean_box(x_247); +x_249 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_249, 0, x_248); +lean_ctor_set(x_249, 1, x_241); +lean_ctor_set(x_8, 1, x_249); +if (lean_is_scalar(x_230)) { + x_250 = lean_alloc_ctor(1, 1, 0); +} else { + x_250 = x_230; +} +lean_ctor_set(x_250, 0, x_8); +x_17 = x_250; +x_18 = x_159; +goto block_23; +} +else +{ +size_t x_251; size_t x_252; lean_object* x_253; uint8_t x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; +x_251 = 0; +x_252 = lean_usize_of_nat(x_237); +lean_dec(x_237); +x_253 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_236, x_251, x_252, x_152); +lean_dec(x_236); +x_254 = 1; +x_255 = lean_box(x_254); +x_256 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_241); +lean_ctor_set(x_8, 1, x_256); +lean_ctor_set(x_8, 0, x_253); +if (lean_is_scalar(x_230)) { + x_257 = lean_alloc_ctor(1, 1, 0); +} else { + x_257 = x_230; +} +lean_ctor_set(x_257, 0, x_8); +x_17 = x_257; +x_18 = x_159; +goto block_23; +} +} +} +} +else +{ +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +lean_dec(x_155); +lean_dec(x_154); +lean_dec(x_153); +lean_free_object(x_8); +lean_dec(x_152); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_258 = lean_ctor_get(x_157, 0); +lean_inc(x_258); +x_259 = lean_ctor_get(x_157, 1); +lean_inc(x_259); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_260 = x_157; +} else { + lean_dec_ref(x_157); + x_260 = lean_box(0); +} +if (lean_is_scalar(x_260)) { + x_261 = lean_alloc_ctor(1, 2, 0); +} else { + x_261 = x_260; +} +lean_ctor_set(x_261, 0, x_258); +lean_ctor_set(x_261, 1, x_259); +return x_261; +} +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; +lean_dec(x_155); +lean_dec(x_16); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + x_262 = x_156; +} else { + lean_dec_ref(x_156); + x_262 = lean_box(0); +} +x_263 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_263, 0, x_153); +lean_ctor_set(x_263, 1, x_154); +lean_ctor_set(x_8, 1, x_263); +if (lean_is_scalar(x_262)) { + x_264 = lean_alloc_ctor(1, 1, 0); +} else { + x_264 = x_262; +} +lean_ctor_set(x_264, 0, x_8); +x_17 = x_264; +x_18 = x_13; +goto block_23; +} +} +} +else +{ +lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +x_265 = lean_ctor_get(x_8, 1); +x_266 = lean_ctor_get(x_8, 0); +lean_inc(x_265); +lean_inc(x_266); +lean_dec(x_8); +x_267 = lean_ctor_get(x_265, 0); +lean_inc(x_267); +x_268 = lean_ctor_get(x_265, 1); +lean_inc(x_268); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_269 = x_265; +} else { + lean_dec_ref(x_265); + x_269 = lean_box(0); +} +x_270 = l_Lean_Expr_fvarId_x21(x_16); +x_271 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_268, x_270); +if (lean_obj_tag(x_271) == 0) +{ +lean_object* x_272; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_272 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_272) == 0) +{ +lean_object* x_273; lean_object* x_274; lean_object* x_275; +x_273 = lean_ctor_get(x_272, 0); +lean_inc(x_273); +x_274 = lean_ctor_get(x_272, 1); +lean_inc(x_274); +lean_dec(x_272); +x_275 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_266, x_270); +if (lean_obj_tag(x_275) == 0) +{ +lean_object* x_276; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_273); +x_276 = l_Lean_Meta_isProp(x_273, x_9, x_10, x_11, x_12, x_274); +if (lean_obj_tag(x_276) == 0) +{ +lean_object* x_277; uint8_t x_278; +x_277 = lean_ctor_get(x_276, 0); +lean_inc(x_277); +x_278 = lean_unbox(x_277); +lean_dec(x_277); +if (x_278 == 0) +{ +lean_object* x_279; lean_object* x_280; +x_279 = lean_ctor_get(x_276, 1); +lean_inc(x_279); +lean_dec(x_276); +lean_inc(x_9); +lean_inc(x_270); +x_280 = l_Lean_FVarId_getDecl(x_270, x_9, x_10, x_11, x_12, x_279); +if (lean_obj_tag(x_280) == 0) +{ +lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; +x_281 = lean_ctor_get(x_280, 0); +lean_inc(x_281); +x_282 = lean_ctor_get(x_280, 1); +lean_inc(x_282); +lean_dec(x_280); +x_283 = l_Lean_LocalDecl_binderInfo(x_281); +lean_dec(x_281); +x_284 = lean_box(x_283); +if (lean_obj_tag(x_284) == 3) +{ +lean_object* x_285; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_273); +lean_inc(x_266); +lean_inc(x_3); +x_285 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_266, x_273, x_9, x_10, x_11, x_12, x_282); +if (lean_obj_tag(x_285) == 0) +{ +lean_object* x_286; uint8_t x_287; +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +x_287 = lean_unbox(x_286); +lean_dec(x_286); +if (x_287 == 0) +{ +lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; +lean_dec(x_273); +lean_dec(x_270); +x_288 = lean_ctor_get(x_285, 1); +lean_inc(x_288); +lean_dec(x_285); +if (lean_is_scalar(x_269)) { + x_289 = lean_alloc_ctor(0, 2, 0); +} else { + x_289 = x_269; +} +lean_ctor_set(x_289, 0, x_267); +lean_ctor_set(x_289, 1, x_268); +x_290 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_290, 0, x_266); +lean_ctor_set(x_290, 1, x_289); +x_291 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_291, 0, x_290); +x_17 = x_291; +x_18 = x_288; +goto block_23; +} +else +{ +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; uint8_t x_303; lean_object* x_304; +lean_dec(x_267); +x_292 = lean_ctor_get(x_285, 1); +lean_inc(x_292); +lean_dec(x_285); +x_293 = lean_box(0); +lean_inc(x_270); +x_294 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_293); +x_295 = lean_box(0); +lean_inc(x_2); +x_296 = lean_array_mk(x_2); +x_297 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_298 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_298, 0, x_297); +lean_ctor_set(x_298, 1, x_295); +lean_ctor_set(x_298, 2, x_296); +x_299 = l_Lean_CollectFVars_main(x_273, x_298); +x_300 = lean_ctor_get(x_299, 2); +lean_inc(x_300); +lean_dec(x_299); +x_301 = lean_array_get_size(x_300); +x_302 = lean_unsigned_to_nat(0u); +x_303 = lean_nat_dec_lt(x_302, x_301); +x_304 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_293); +if (x_303 == 0) +{ +uint8_t x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; +lean_dec(x_301); +lean_dec(x_300); +x_305 = 1; +x_306 = lean_box(x_305); +if (lean_is_scalar(x_269)) { + x_307 = lean_alloc_ctor(0, 2, 0); +} else { + x_307 = x_269; +} +lean_ctor_set(x_307, 0, x_306); +lean_ctor_set(x_307, 1, x_304); +x_308 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_308, 0, x_294); +lean_ctor_set(x_308, 1, x_307); +x_309 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_309, 0, x_308); +x_17 = x_309; +x_18 = x_292; +goto block_23; +} +else +{ +uint8_t x_310; +x_310 = lean_nat_dec_le(x_301, x_301); +if (x_310 == 0) +{ +uint8_t x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; +lean_dec(x_301); +lean_dec(x_300); +x_311 = 1; +x_312 = lean_box(x_311); +if (lean_is_scalar(x_269)) { + x_313 = lean_alloc_ctor(0, 2, 0); +} else { + x_313 = x_269; +} +lean_ctor_set(x_313, 0, x_312); +lean_ctor_set(x_313, 1, x_304); +x_314 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_314, 0, x_294); +lean_ctor_set(x_314, 1, x_313); +x_315 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_315, 0, x_314); +x_17 = x_315; +x_18 = x_292; +goto block_23; +} +else +{ +size_t x_316; size_t x_317; lean_object* x_318; uint8_t x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; +x_316 = 0; +x_317 = lean_usize_of_nat(x_301); +lean_dec(x_301); +x_318 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_300, x_316, x_317, x_294); +lean_dec(x_300); +x_319 = 1; +x_320 = lean_box(x_319); +if (lean_is_scalar(x_269)) { + x_321 = lean_alloc_ctor(0, 2, 0); +} else { + x_321 = x_269; +} +lean_ctor_set(x_321, 0, x_320); +lean_ctor_set(x_321, 1, x_304); +x_322 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_322, 0, x_318); +lean_ctor_set(x_322, 1, x_321); +x_323 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_323, 0, x_322); +x_17 = x_323; +x_18 = x_292; +goto block_23; +} +} +} +} +else +{ +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_324 = lean_ctor_get(x_285, 0); +lean_inc(x_324); +x_325 = lean_ctor_get(x_285, 1); +lean_inc(x_325); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_326 = x_285; +} else { + lean_dec_ref(x_285); + x_326 = lean_box(0); +} +if (lean_is_scalar(x_326)) { + x_327 = lean_alloc_ctor(1, 2, 0); +} else { + x_327 = x_326; +} +lean_ctor_set(x_327, 0, x_324); +lean_ctor_set(x_327, 1, x_325); +return x_327; +} +} +else +{ +lean_object* x_328; lean_object* x_329; lean_object* x_330; +lean_dec(x_284); +lean_dec(x_273); +lean_dec(x_270); +if (lean_is_scalar(x_269)) { + x_328 = lean_alloc_ctor(0, 2, 0); +} else { + x_328 = x_269; +} +lean_ctor_set(x_328, 0, x_267); +lean_ctor_set(x_328, 1, x_268); +x_329 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_329, 0, x_266); +lean_ctor_set(x_329, 1, x_328); +x_330 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_330, 0, x_329); +x_17 = x_330; +x_18 = x_282; +goto block_23; +} +} +else +{ +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_331 = lean_ctor_get(x_280, 0); +lean_inc(x_331); +x_332 = lean_ctor_get(x_280, 1); +lean_inc(x_332); +if (lean_is_exclusive(x_280)) { + lean_ctor_release(x_280, 0); + lean_ctor_release(x_280, 1); + x_333 = x_280; +} else { + lean_dec_ref(x_280); + x_333 = lean_box(0); +} +if (lean_is_scalar(x_333)) { + x_334 = lean_alloc_ctor(1, 2, 0); +} else { + x_334 = x_333; +} +lean_ctor_set(x_334, 0, x_331); +lean_ctor_set(x_334, 1, x_332); +return x_334; +} +} +else +{ +lean_object* x_335; uint8_t x_336; +x_335 = lean_ctor_get(x_276, 1); +lean_inc(x_335); +lean_dec(x_276); +x_336 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_266, x_273); +if (x_336 == 0) +{ +lean_object* x_337; lean_object* x_338; lean_object* x_339; +lean_dec(x_270); +if (lean_is_scalar(x_269)) { + x_337 = lean_alloc_ctor(0, 2, 0); +} else { + x_337 = x_269; +} +lean_ctor_set(x_337, 0, x_267); +lean_ctor_set(x_337, 1, x_268); +x_338 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_338, 0, x_266); +lean_ctor_set(x_338, 1, x_337); +x_339 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_339, 0, x_338); +x_17 = x_339; +x_18 = x_335; +goto block_23; +} +else +{ +lean_object* x_340; lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; +lean_dec(x_267); +x_340 = lean_box(0); +lean_inc(x_270); +x_341 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_340); +x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_340); +x_343 = 1; +x_344 = lean_box(x_343); +if (lean_is_scalar(x_269)) { + x_345 = lean_alloc_ctor(0, 2, 0); +} else { + x_345 = x_269; +} +lean_ctor_set(x_345, 0, x_344); +lean_ctor_set(x_345, 1, x_342); +x_346 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_346, 0, x_341); +lean_ctor_set(x_346, 1, x_345); +x_347 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_347, 0, x_346); +x_17 = x_347; +x_18 = x_335; +goto block_23; +} +} +} +else +{ +lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; +lean_dec(x_273); +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_348 = lean_ctor_get(x_276, 0); +lean_inc(x_348); +x_349 = lean_ctor_get(x_276, 1); +lean_inc(x_349); +if (lean_is_exclusive(x_276)) { + lean_ctor_release(x_276, 0); + lean_ctor_release(x_276, 1); + x_350 = x_276; +} else { + lean_dec_ref(x_276); + x_350 = lean_box(0); +} +if (lean_is_scalar(x_350)) { + x_351 = lean_alloc_ctor(1, 2, 0); +} else { + x_351 = x_350; +} +lean_ctor_set(x_351, 0, x_348); +lean_ctor_set(x_351, 1, x_349); +return x_351; +} +} +else +{ +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; uint8_t x_361; lean_object* x_362; lean_object* x_363; +lean_dec(x_267); +if (lean_is_exclusive(x_275)) { + lean_ctor_release(x_275, 0); + x_352 = x_275; +} else { + lean_dec_ref(x_275); + x_352 = lean_box(0); +} +x_353 = lean_box(0); +lean_inc(x_2); +x_354 = lean_array_mk(x_2); +x_355 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; +x_356 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_356, 0, x_355); +lean_ctor_set(x_356, 1, x_353); +lean_ctor_set(x_356, 2, x_354); +x_357 = l_Lean_CollectFVars_main(x_273, x_356); +x_358 = lean_ctor_get(x_357, 2); +lean_inc(x_358); +lean_dec(x_357); +x_359 = lean_array_get_size(x_358); +x_360 = lean_unsigned_to_nat(0u); +x_361 = lean_nat_dec_lt(x_360, x_359); +x_362 = lean_box(0); +x_363 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_362); +if (x_361 == 0) +{ +uint8_t x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; +lean_dec(x_359); +lean_dec(x_358); +x_364 = 1; +x_365 = lean_box(x_364); +if (lean_is_scalar(x_269)) { + x_366 = lean_alloc_ctor(0, 2, 0); +} else { + x_366 = x_269; +} +lean_ctor_set(x_366, 0, x_365); +lean_ctor_set(x_366, 1, x_363); +x_367 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_367, 0, x_266); +lean_ctor_set(x_367, 1, x_366); +if (lean_is_scalar(x_352)) { + x_368 = lean_alloc_ctor(1, 1, 0); +} else { + x_368 = x_352; +} +lean_ctor_set(x_368, 0, x_367); +x_17 = x_368; +x_18 = x_274; +goto block_23; +} +else +{ +uint8_t x_369; +x_369 = lean_nat_dec_le(x_359, x_359); +if (x_369 == 0) +{ +uint8_t x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; +lean_dec(x_359); +lean_dec(x_358); +x_370 = 1; +x_371 = lean_box(x_370); +if (lean_is_scalar(x_269)) { + x_372 = lean_alloc_ctor(0, 2, 0); +} else { + x_372 = x_269; +} +lean_ctor_set(x_372, 0, x_371); +lean_ctor_set(x_372, 1, x_363); +x_373 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_373, 0, x_266); +lean_ctor_set(x_373, 1, x_372); +if (lean_is_scalar(x_352)) { + x_374 = lean_alloc_ctor(1, 1, 0); +} else { + x_374 = x_352; +} +lean_ctor_set(x_374, 0, x_373); +x_17 = x_374; +x_18 = x_274; +goto block_23; +} +else +{ +size_t x_375; size_t x_376; lean_object* x_377; uint8_t x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; +x_375 = 0; +x_376 = lean_usize_of_nat(x_359); +lean_dec(x_359); +x_377 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_358, x_375, x_376, x_266); +lean_dec(x_358); +x_378 = 1; +x_379 = lean_box(x_378); +if (lean_is_scalar(x_269)) { + x_380 = lean_alloc_ctor(0, 2, 0); +} else { + x_380 = x_269; +} +lean_ctor_set(x_380, 0, x_379); +lean_ctor_set(x_380, 1, x_363); +x_381 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_381, 0, x_377); +lean_ctor_set(x_381, 1, x_380); +if (lean_is_scalar(x_352)) { + x_382 = lean_alloc_ctor(1, 1, 0); +} else { + x_382 = x_352; +} +lean_ctor_set(x_382, 0, x_381); +x_17 = x_382; +x_18 = x_274; +goto block_23; +} +} +} +} +else +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +x_383 = lean_ctor_get(x_272, 0); +lean_inc(x_383); +x_384 = lean_ctor_get(x_272, 1); +lean_inc(x_384); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + lean_ctor_release(x_272, 1); + x_385 = x_272; +} else { + lean_dec_ref(x_272); + x_385 = lean_box(0); +} +if (lean_is_scalar(x_385)) { + x_386 = lean_alloc_ctor(1, 2, 0); +} else { + x_386 = x_385; +} +lean_ctor_set(x_386, 0, x_383); +lean_ctor_set(x_386, 1, x_384); +return x_386; +} +} +else +{ +lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; +lean_dec(x_270); +lean_dec(x_16); +if (lean_is_exclusive(x_271)) { + lean_ctor_release(x_271, 0); + x_387 = x_271; +} else { + lean_dec_ref(x_271); + x_387 = lean_box(0); +} +if (lean_is_scalar(x_269)) { + x_388 = lean_alloc_ctor(0, 2, 0); +} else { + x_388 = x_269; +} +lean_ctor_set(x_388, 0, x_267); +lean_ctor_set(x_388, 1, x_268); +x_389 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_389, 0, x_266); +lean_ctor_set(x_389, 1, x_388); +if (lean_is_scalar(x_387)) { + x_390 = lean_alloc_ctor(1, 1, 0); +} else { + x_390 = x_387; +} +lean_ctor_set(x_390, 0, x_389); +x_17 = x_390; +x_18 = x_13; +goto block_23; +} +} +block_23: +{ +lean_object* x_19; size_t x_20; size_t x_21; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = 1; +x_21 = lean_usize_add(x_7, x_20); +x_7 = x_21; +x_8 = x_19; +x_13 = x_18; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_box(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_3); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_10); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_box(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_2); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_3); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_4); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_10); +return x_22; +} +} +} +static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_13, 1); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_15, 0); +lean_dec(x_17); +x_18 = lean_box(0); +x_19 = 0; +x_20 = lean_box(x_19); +lean_ctor_set(x_15, 0, x_20); +x_21 = lean_array_size(x_2); +x_22 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_3); +x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_18, x_2, x_21, x_22, x_13, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_23, 1); +x_28 = lean_ctor_get(x_23, 0); +lean_dec(x_28); +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_24, 0); +x_31 = lean_ctor_get(x_24, 1); +lean_dec(x_31); +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_25, 0); +x_34 = lean_ctor_get(x_25, 1); +x_35 = lean_unsigned_to_nat(0u); +x_36 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_35, x_30); +x_37 = lean_nat_dec_eq(x_36, x_1); +lean_dec(x_36); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; +lean_free_object(x_25); +lean_free_object(x_24); +lean_free_object(x_23); +x_38 = lean_box(0); +x_39 = lean_unbox(x_33); +lean_dec(x_33); +lean_inc(x_6); +x_40 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_39, x_34, x_30, x_6, x_38, x_8, x_9, x_10, x_11, x_27); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +if (lean_obj_tag(x_41) == 0) +{ +uint8_t x_42; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_42 = !lean_is_exclusive(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 0); +lean_dec(x_43); +x_44 = lean_ctor_get(x_41, 0); +lean_inc(x_44); +lean_dec(x_41); +lean_ctor_set(x_40, 0, x_44); +return x_40; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_ctor_get(x_41, 0); +lean_inc(x_46); +lean_dec(x_41); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_40, 1); +lean_inc(x_48); +lean_dec(x_40); +x_49 = lean_ctor_get(x_41, 0); +lean_inc(x_49); +lean_dec(x_41); +x_7 = x_49; +x_12 = x_48; +goto _start; +} +} +else +{ +lean_object* x_51; lean_object* x_52; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_51 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_24); +lean_ctor_set(x_23, 0, x_52); +return x_23; +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_53 = lean_ctor_get(x_25, 0); +x_54 = lean_ctor_get(x_25, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_25); +x_55 = lean_unsigned_to_nat(0u); +x_56 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_55, x_30); +x_57 = lean_nat_dec_eq(x_56, x_1); +lean_dec(x_56); +if (x_57 == 0) +{ +lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; +lean_free_object(x_24); +lean_free_object(x_23); +x_58 = lean_box(0); +x_59 = lean_unbox(x_53); +lean_dec(x_53); +lean_inc(x_6); +x_60 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_59, x_54, x_30, x_6, x_58, x_8, x_9, x_10, x_11, x_27); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_63 = x_60; +} else { + lean_dec_ref(x_60); + x_63 = lean_box(0); +} +x_64 = lean_ctor_get(x_61, 0); +lean_inc(x_64); +lean_dec(x_61); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_63; +} +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; +} +else +{ +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +lean_dec(x_60); +x_67 = lean_ctor_get(x_61, 0); +lean_inc(x_67); +lean_dec(x_61); +x_7 = x_67; +x_12 = x_66; +goto _start; +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_53); +lean_ctor_set(x_69, 1, x_54); +lean_ctor_set(x_24, 1, x_69); +x_70 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_24); +lean_ctor_set(x_23, 0, x_71); +return x_23; +} +} +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_72 = lean_ctor_get(x_24, 0); +lean_inc(x_72); +lean_dec(x_24); +x_73 = lean_ctor_get(x_25, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_25, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_25)) { + lean_ctor_release(x_25, 0); + lean_ctor_release(x_25, 1); + x_75 = x_25; +} else { + lean_dec_ref(x_25); + x_75 = lean_box(0); +} +x_76 = lean_unsigned_to_nat(0u); +x_77 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_76, x_72); +x_78 = lean_nat_dec_eq(x_77, x_1); +lean_dec(x_77); +if (x_78 == 0) +{ +lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_75); +lean_free_object(x_23); +x_79 = lean_box(0); +x_80 = lean_unbox(x_73); +lean_dec(x_73); +lean_inc(x_6); +x_81 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_80, x_74, x_72, x_6, x_79, x_8, x_9, x_10, x_11, x_27); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +if (lean_obj_tag(x_82) == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_84 = x_81; +} else { + lean_dec_ref(x_81); + x_84 = lean_box(0); +} +x_85 = lean_ctor_get(x_82, 0); +lean_inc(x_85); +lean_dec(x_82); +if (lean_is_scalar(x_84)) { + x_86 = lean_alloc_ctor(0, 2, 0); +} else { + x_86 = x_84; +} +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_83); +return x_86; +} +else +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_81, 1); +lean_inc(x_87); +lean_dec(x_81); +x_88 = lean_ctor_get(x_82, 0); +lean_inc(x_88); +lean_dec(x_82); +x_7 = x_88; +x_12 = x_87; +goto _start; +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +if (lean_is_scalar(x_75)) { + x_90 = lean_alloc_ctor(0, 2, 0); +} else { + x_90 = x_75; +} +lean_ctor_set(x_90, 0, x_73); +lean_ctor_set(x_90, 1, x_74); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_72); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_91); +lean_ctor_set(x_23, 0, x_93); +return x_23; +} +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; +x_94 = lean_ctor_get(x_23, 1); +lean_inc(x_94); +lean_dec(x_23); +x_95 = lean_ctor_get(x_24, 0); +lean_inc(x_95); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_96 = x_24; +} else { + lean_dec_ref(x_24); + x_96 = lean_box(0); +} +x_97 = lean_ctor_get(x_25, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_25, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_25)) { + lean_ctor_release(x_25, 0); + lean_ctor_release(x_25, 1); + x_99 = x_25; +} else { + lean_dec_ref(x_25); + x_99 = lean_box(0); +} +x_100 = lean_unsigned_to_nat(0u); +x_101 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_100, x_95); +x_102 = lean_nat_dec_eq(x_101, x_1); +lean_dec(x_101); +if (x_102 == 0) +{ +lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_99); +lean_dec(x_96); +x_103 = lean_box(0); +x_104 = lean_unbox(x_97); +lean_dec(x_97); +lean_inc(x_6); +x_105 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_104, x_98, x_95, x_6, x_103, x_8, x_9, x_10, x_11, x_94); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_108 = x_105; +} else { + lean_dec_ref(x_105); + x_108 = lean_box(0); +} +x_109 = lean_ctor_get(x_106, 0); +lean_inc(x_109); +lean_dec(x_106); +if (lean_is_scalar(x_108)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_108; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_107); +return x_110; +} +else +{ +lean_object* x_111; lean_object* x_112; +x_111 = lean_ctor_get(x_105, 1); +lean_inc(x_111); +lean_dec(x_105); +x_112 = lean_ctor_get(x_106, 0); +lean_inc(x_112); +lean_dec(x_106); +x_7 = x_112; +x_12 = x_111; +goto _start; +} +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +if (lean_is_scalar(x_99)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_99; +} +lean_ctor_set(x_114, 0, x_97); +lean_ctor_set(x_114, 1, x_98); +if (lean_is_scalar(x_96)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_96; +} +lean_ctor_set(x_115, 0, x_95); +lean_ctor_set(x_115, 1, x_114); +x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_115); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_94); +return x_118; +} +} +} +else +{ +uint8_t x_119; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_119 = !lean_is_exclusive(x_23); +if (x_119 == 0) +{ +return x_23; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_23, 0); +x_121 = lean_ctor_get(x_23, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_23); +x_122 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +return x_122; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; uint8_t x_125; lean_object* x_126; lean_object* x_127; size_t x_128; size_t x_129; lean_object* x_130; +x_123 = lean_ctor_get(x_15, 1); +lean_inc(x_123); +lean_dec(x_15); +x_124 = lean_box(0); +x_125 = 0; +x_126 = lean_box(x_125); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_123); +lean_ctor_set(x_13, 1, x_127); +x_128 = lean_array_size(x_2); +x_129 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_3); +x_130 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_124, x_2, x_128, x_129, x_13, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_131, 1); +lean_inc(x_132); +x_133 = lean_ctor_get(x_130, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_134 = x_130; +} else { + lean_dec_ref(x_130); + x_134 = lean_box(0); +} +x_135 = lean_ctor_get(x_131, 0); +lean_inc(x_135); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_136 = x_131; +} else { + lean_dec_ref(x_131); + x_136 = lean_box(0); +} +x_137 = lean_ctor_get(x_132, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_132, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_139 = x_132; +} else { + lean_dec_ref(x_132); + x_139 = lean_box(0); +} +x_140 = lean_unsigned_to_nat(0u); +x_141 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_140, x_135); +x_142 = lean_nat_dec_eq(x_141, x_1); +lean_dec(x_141); +if (x_142 == 0) +{ +lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_139); +lean_dec(x_136); +lean_dec(x_134); +x_143 = lean_box(0); +x_144 = lean_unbox(x_137); +lean_dec(x_137); +lean_inc(x_6); +x_145 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_144, x_138, x_135, x_6, x_143, x_8, x_9, x_10, x_11, x_133); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +if (lean_obj_tag(x_146) == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_148 = x_145; +} else { + lean_dec_ref(x_145); + x_148 = lean_box(0); +} +x_149 = lean_ctor_get(x_146, 0); +lean_inc(x_149); +lean_dec(x_146); +if (lean_is_scalar(x_148)) { + x_150 = lean_alloc_ctor(0, 2, 0); +} else { + x_150 = x_148; +} +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_147); +return x_150; +} +else +{ +lean_object* x_151; lean_object* x_152; +x_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +lean_dec(x_145); +x_152 = lean_ctor_get(x_146, 0); +lean_inc(x_152); +lean_dec(x_146); +x_7 = x_152; +x_12 = x_151; +goto _start; +} +} +else +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +if (lean_is_scalar(x_139)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_139; +} +lean_ctor_set(x_154, 0, x_137); +lean_ctor_set(x_154, 1, x_138); +if (lean_is_scalar(x_136)) { + x_155 = lean_alloc_ctor(0, 2, 0); +} else { + x_155 = x_136; +} +lean_ctor_set(x_155, 0, x_135); +lean_ctor_set(x_155, 1, x_154); +x_156 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_157 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_155); +if (lean_is_scalar(x_134)) { + x_158 = lean_alloc_ctor(0, 2, 0); +} else { + x_158 = x_134; +} +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_133); +return x_158; +} +} +else +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_159 = lean_ctor_get(x_130, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_130, 1); +lean_inc(x_160); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_161 = x_130; +} else { + lean_dec_ref(x_130); + x_161 = lean_box(0); +} +if (lean_is_scalar(x_161)) { + x_162 = lean_alloc_ctor(1, 2, 0); +} else { + x_162 = x_161; +} +lean_ctor_set(x_162, 0, x_159); +lean_ctor_set(x_162, 1, x_160); +return x_162; +} +} +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; size_t x_172; size_t x_173; lean_object* x_174; +x_163 = lean_ctor_get(x_13, 1); +x_164 = lean_ctor_get(x_13, 0); +lean_inc(x_163); +lean_inc(x_164); +lean_dec(x_13); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +if (lean_is_exclusive(x_163)) { + lean_ctor_release(x_163, 0); + lean_ctor_release(x_163, 1); + x_166 = x_163; +} else { + lean_dec_ref(x_163); + x_166 = lean_box(0); +} +x_167 = lean_box(0); +x_168 = 0; +x_169 = lean_box(x_168); +if (lean_is_scalar(x_166)) { + x_170 = lean_alloc_ctor(0, 2, 0); +} else { + x_170 = x_166; +} +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_165); +x_171 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_171, 0, x_164); +lean_ctor_set(x_171, 1, x_170); +x_172 = lean_array_size(x_2); +x_173 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_3); +x_174 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_167, x_2, x_172, x_173, x_171, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_175, 1); +lean_inc(x_176); +x_177 = lean_ctor_get(x_174, 1); +lean_inc(x_177); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_178 = x_174; +} else { + lean_dec_ref(x_174); + x_178 = lean_box(0); +} +x_179 = lean_ctor_get(x_175, 0); +lean_inc(x_179); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_180 = x_175; +} else { + lean_dec_ref(x_175); + x_180 = lean_box(0); +} +x_181 = lean_ctor_get(x_176, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_176, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_176)) { + lean_ctor_release(x_176, 0); + lean_ctor_release(x_176, 1); + x_183 = x_176; +} else { + lean_dec_ref(x_176); + x_183 = lean_box(0); +} +x_184 = lean_unsigned_to_nat(0u); +x_185 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_184, x_179); +x_186 = lean_nat_dec_eq(x_185, x_1); +lean_dec(x_185); +if (x_186 == 0) +{ +lean_object* x_187; uint8_t x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_183); +lean_dec(x_180); +lean_dec(x_178); +x_187 = lean_box(0); +x_188 = lean_unbox(x_181); +lean_dec(x_181); +lean_inc(x_6); +x_189 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_188, x_182, x_179, x_6, x_187, x_8, x_9, x_10, x_11, x_177); +x_190 = lean_ctor_get(x_189, 0); +lean_inc(x_190); +if (lean_obj_tag(x_190) == 0) +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_10; +lean_dec(x_3); +x_191 = lean_ctor_get(x_189, 1); +lean_inc(x_191); +if (lean_is_exclusive(x_189)) { + lean_ctor_release(x_189, 0); + lean_ctor_release(x_189, 1); + x_192 = x_189; +} else { + lean_dec_ref(x_189); + x_192 = lean_box(0); } +x_193 = lean_ctor_get(x_190, 0); +lean_inc(x_193); +lean_dec(x_190); +if (lean_is_scalar(x_192)) { + x_194 = lean_alloc_ctor(0, 2, 0); +} else { + x_194 = x_192; } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_ctor_set(x_194, 0, x_193); +lean_ctor_set(x_194, 1, x_191); +return x_194; +} +else { -lean_object* x_14; -x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_14; +lean_object* x_195; lean_object* x_196; +x_195 = lean_ctor_get(x_189, 1); +lean_inc(x_195); +lean_dec(x_189); +x_196 = lean_ctor_get(x_190, 0); +lean_inc(x_196); +lean_dec(x_190); +x_7 = x_196; +x_12 = x_195; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; -x_8 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_8; -} +if (lean_is_scalar(x_183)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_183; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_2); -return x_9; +lean_ctor_set(x_198, 0, x_181); +lean_ctor_set(x_198, 1, x_182); +if (lean_is_scalar(x_180)) { + x_199 = lean_alloc_ctor(0, 2, 0); +} else { + x_199 = x_180; } +lean_ctor_set(x_199, 0, x_179); +lean_ctor_set(x_199, 1, x_198); +x_200 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_201 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_199); +if (lean_is_scalar(x_178)) { + x_202 = lean_alloc_ctor(0, 2, 0); +} else { + x_202 = x_178; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_2); -lean_dec(x_2); -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); -return x_10; +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_177); +return x_202; } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -if (lean_obj_tag(x_1) == 0) +else { -lean_object* x_9; lean_object* x_10; -lean_dec(x_7); +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -x_9 = l_List_reverse___rarg(x_2); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; +x_203 = lean_ctor_get(x_174, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_174, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_205 = x_174; +} else { + lean_dec_ref(x_174); + x_205 = lean_box(0); } -else +if (lean_is_scalar(x_205)) { + x_206 = lean_alloc_ctor(1, 2, 0); +} else { + x_206 = x_205; +} +lean_ctor_set(x_206, 0, x_203); +lean_ctor_set(x_206, 1, x_204); +return x_206; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_1); -if (x_11 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_6, 1); +lean_inc(x_12); +lean_dec(x_6); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_1, 0); -x_13 = lean_ctor_get(x_1, 1); -x_14 = 0; +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 1); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_14, 0); +lean_dec(x_16); +x_17 = lean_box(0); +x_18 = 0; +x_19 = lean_box(x_18); +lean_ctor_set(x_14, 0, x_19); +x_20 = lean_array_size(x_2); +x_21 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_12, x_14, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_15) == 0) +x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_17, x_2, x_20, x_21, x_12, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_16); +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +x_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) { -lean_object* _tmp_0 = x_13; -lean_object* _tmp_1 = x_1; -lean_object* _tmp_7 = x_17; -x_1 = _tmp_0; -x_2 = _tmp_1; -x_8 = _tmp_7; -} -goto _start; -} -else +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_22, 1); +x_27 = lean_ctor_get(x_22, 0); +lean_dec(x_27); +x_28 = !lean_is_exclusive(x_23); +if (x_28 == 0) { -uint8_t x_19; -lean_free_object(x_1); -lean_dec(x_13); +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_23, 0); +x_30 = lean_ctor_get(x_23, 1); +lean_dec(x_30); +x_31 = !lean_is_exclusive(x_24); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_24, 0); +x_33 = lean_ctor_get(x_24, 1); +x_34 = lean_unsigned_to_nat(0u); +x_35 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_34, x_29); +x_36 = lean_nat_dec_eq(x_35, x_1); +lean_dec(x_35); +if (x_36 == 0) +{ +lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; +lean_free_object(x_24); +lean_free_object(x_23); +lean_free_object(x_22); +x_37 = lean_box(0); +x_38 = lean_unbox(x_32); +lean_dec(x_32); +lean_inc(x_5); +x_39 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_38, x_33, x_29, x_5, x_37, x_7, x_8, x_9, x_10, x_26); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_obj_tag(x_40) == 0) +{ +uint8_t x_41; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_19 = !lean_is_exclusive(x_15); -if (x_19 == 0) +x_41 = !lean_is_exclusive(x_39); +if (x_41 == 0) { -return x_15; +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_39, 0); +lean_dec(x_42); +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +lean_dec(x_40); +lean_ctor_set(x_39, 0, x_43); +return x_39; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_15, 0); -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_15); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_39, 1); +lean_inc(x_44); +lean_dec(x_39); +x_45 = lean_ctor_get(x_40, 0); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +return x_46; } } else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 0); -x_24 = lean_ctor_get(x_1, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_1); -x_25 = 0; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_23, x_25, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_2); -x_1 = x_24; -x_2 = x_29; -x_8 = x_28; +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_dec(x_39); +x_48 = lean_ctor_get(x_40, 0); +lean_inc(x_48); +lean_dec(x_40); +x_6 = x_48; +x_11 = x_47; goto _start; } +} else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_24); +lean_object* x_50; lean_object* x_51; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_33 = x_26; -} else { - lean_dec_ref(x_26); - x_33 = lean_box(0); -} -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(1, 2, 0); -} else { - x_34 = x_33; -} -lean_ctor_set(x_34, 0, x_31); -lean_ctor_set(x_34, 1, x_32); -return x_34; -} -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; +x_50 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_23); +lean_ctor_set(x_22, 0, x_51); +return x_22; } } -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(10u); -x_2 = lean_unsigned_to_nat(1u); -x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__3() { -_start: +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_24, 0); +x_53 = lean_ctor_get(x_24, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_24); +x_54 = lean_unsigned_to_nat(0u); +x_55 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_54, x_29); +x_56 = lean_nat_dec_eq(x_55, x_1); +lean_dec(x_55); +if (x_56 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__2; -x_3 = lean_mk_array(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__4() { -_start: +lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; +lean_free_object(x_23); +lean_free_object(x_22); +x_57 = lean_box(0); +x_58 = lean_unbox(x_52); +lean_dec(x_52); +lean_inc(x_5); +x_59 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_58, x_53, x_29, x_5, x_57, x_7, x_8, x_9, x_10, x_26); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__3; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_62 = x_59; +} else { + lean_dec_ref(x_59); + x_62 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_Grind_NormalizePattern_main___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; -x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_3 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 2, x_2); -return x_3; +x_63 = lean_ctor_get(x_60, 0); +lean_inc(x_63); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_62; } +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_61); +return x_64; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_NormalizePattern_main(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_7 = lean_box(0); -x_8 = l_Lean_Meta_Grind_NormalizePattern_main___closed__5; -x_9 = lean_st_mk_ref(x_8, x_6); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -x_13 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_11, x_2, x_3, x_4, x_5, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_st_ref_get(x_11, x_15); -lean_dec(x_11); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_array_to_list(x_19); -x_21 = lean_ctor_get(x_18, 2); -lean_inc(x_21); -lean_dec(x_18); -lean_ctor_set(x_9, 1, x_21); -lean_ctor_set(x_9, 0, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_14); -lean_ctor_set(x_22, 1, x_9); -lean_ctor_set(x_16, 0, x_22); -return x_16; +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_59, 1); +lean_inc(x_65); +lean_dec(x_59); +x_66 = lean_ctor_get(x_60, 0); +lean_inc(x_66); +lean_dec(x_60); +x_6 = x_66; +x_11 = x_65; +goto _start; +} } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_23 = lean_ctor_get(x_16, 0); -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_16); -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -x_26 = lean_array_to_list(x_25); -x_27 = lean_ctor_get(x_23, 2); -lean_inc(x_27); -lean_dec(x_23); -lean_ctor_set(x_9, 1, x_27); -lean_ctor_set(x_9, 0, x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_14); -lean_ctor_set(x_28, 1, x_9); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_24); -return x_29; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_52); +lean_ctor_set(x_68, 1, x_53); +lean_ctor_set(x_23, 1, x_68); +x_69 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_23); +lean_ctor_set(x_22, 0, x_70); +return x_22; +} } } else { -uint8_t x_30; -lean_free_object(x_9); -lean_dec(x_11); -x_30 = !lean_is_exclusive(x_13); -if (x_30 == 0) +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_71 = lean_ctor_get(x_23, 0); +lean_inc(x_71); +lean_dec(x_23); +x_72 = lean_ctor_get(x_24, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_24, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_74 = x_24; +} else { + lean_dec_ref(x_24); + x_74 = lean_box(0); +} +x_75 = lean_unsigned_to_nat(0u); +x_76 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_75, x_71); +x_77 = lean_nat_dec_eq(x_76, x_1); +lean_dec(x_76); +if (x_77 == 0) { -return x_13; +lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_74); +lean_free_object(x_22); +x_78 = lean_box(0); +x_79 = lean_unbox(x_72); +lean_dec(x_72); +lean_inc(x_5); +x_80 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_79, x_73, x_71, x_5, x_78, x_7, x_8, x_9, x_10, x_26); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (lean_obj_tag(x_81) == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_83 = x_80; +} else { + lean_dec_ref(x_80); + x_83 = lean_box(0); +} +x_84 = lean_ctor_get(x_81, 0); +lean_inc(x_84); +lean_dec(x_81); +if (lean_is_scalar(x_83)) { + x_85 = lean_alloc_ctor(0, 2, 0); +} else { + x_85 = x_83; +} +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +return x_85; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_13, 0); -x_32 = lean_ctor_get(x_13, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_13); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} +lean_object* x_86; lean_object* x_87; +x_86 = lean_ctor_get(x_80, 1); +lean_inc(x_86); +lean_dec(x_80); +x_87 = lean_ctor_get(x_81, 0); +lean_inc(x_87); +lean_dec(x_81); +x_6 = x_87; +x_11 = x_86; +goto _start; } } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_9, 0); -x_35 = lean_ctor_get(x_9, 1); -lean_inc(x_35); -lean_inc(x_34); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_10); lean_dec(x_9); -lean_inc(x_34); -x_36 = l_List_mapM_loop___at_Lean_Meta_Grind_NormalizePattern_main___spec__1(x_1, x_7, x_34, x_2, x_3, x_4, x_5, x_35); -if (lean_obj_tag(x_36) == 0) +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +if (lean_is_scalar(x_74)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_74; +} +lean_ctor_set(x_89, 0, x_72); +lean_ctor_set(x_89, 1, x_73); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_71); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_90); +lean_ctor_set(x_22, 0, x_92); +return x_22; +} +} +} +else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_st_ref_get(x_34, x_38); -lean_dec(x_34); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_42 = x_39; +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_93 = lean_ctor_get(x_22, 1); +lean_inc(x_93); +lean_dec(x_22); +x_94 = lean_ctor_get(x_23, 0); +lean_inc(x_94); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + lean_ctor_release(x_23, 1); + x_95 = x_23; } else { - lean_dec_ref(x_39); - x_42 = lean_box(0); + lean_dec_ref(x_23); + x_95 = lean_box(0); } -x_43 = lean_ctor_get(x_40, 0); -lean_inc(x_43); -x_44 = lean_array_to_list(x_43); -x_45 = lean_ctor_get(x_40, 2); -lean_inc(x_45); -lean_dec(x_40); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_37); -lean_ctor_set(x_47, 1, x_46); -if (lean_is_scalar(x_42)) { - x_48 = lean_alloc_ctor(0, 2, 0); +x_96 = lean_ctor_get(x_24, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_24, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_98 = x_24; +} else { + lean_dec_ref(x_24); + x_98 = lean_box(0); +} +x_99 = lean_unsigned_to_nat(0u); +x_100 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_99, x_94); +x_101 = lean_nat_dec_eq(x_100, x_1); +lean_dec(x_100); +if (x_101 == 0) +{ +lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_98); +lean_dec(x_95); +x_102 = lean_box(0); +x_103 = lean_unbox(x_96); +lean_dec(x_96); +lean_inc(x_5); +x_104 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_103, x_97, x_94, x_5, x_102, x_7, x_8, x_9, x_10, x_93); +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_107 = x_104; } else { - x_48 = x_42; + lean_dec_ref(x_104); + x_107 = lean_box(0); } -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_41); -return x_48; +x_108 = lean_ctor_get(x_105, 0); +lean_inc(x_108); +lean_dec(x_105); +if (lean_is_scalar(x_107)) { + x_109 = lean_alloc_ctor(0, 2, 0); +} else { + x_109 = x_107; +} +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_106); +return x_109; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -lean_dec(x_34); -x_49 = lean_ctor_get(x_36, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_36, 1); -lean_inc(x_50); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_51 = x_36; -} else { - lean_dec_ref(x_36); - x_51 = lean_box(0); +lean_object* x_110; lean_object* x_111; +x_110 = lean_ctor_get(x_104, 1); +lean_inc(x_110); +lean_dec(x_104); +x_111 = lean_ctor_get(x_105, 0); +lean_inc(x_111); +lean_dec(x_105); +x_6 = x_111; +x_11 = x_110; +goto _start; } -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(1, 2, 0); -} else { - x_52 = x_51; } -lean_ctor_set(x_52, 0, x_49); -lean_ctor_set(x_52, 1, x_50); -return x_52; +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +if (lean_is_scalar(x_98)) { + x_113 = lean_alloc_ctor(0, 2, 0); +} else { + x_113 = x_98; } +lean_ctor_set(x_113, 0, x_96); +lean_ctor_set(x_113, 1, x_97); +if (lean_is_scalar(x_95)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_95; } +lean_ctor_set(x_114, 0, x_94); +lean_ctor_set(x_114, 1, x_113); +x_115 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_114); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_93); +return x_117; } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_5, x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_array_uget(x_4, x_5); -x_9 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_8); -if (lean_obj_tag(x_9) == 0) -{ -size_t x_10; size_t x_11; -lean_dec(x_8); -x_10 = 1; -x_11 = lean_usize_add(x_5, x_10); -x_5 = x_11; -goto _start; } else { -lean_object* x_13; +uint8_t x_118; +lean_dec(x_10); lean_dec(x_9); -x_13 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); lean_dec(x_8); -if (lean_obj_tag(x_13) == 0) +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_118 = !lean_is_exclusive(x_22); +if (x_118 == 0) { -uint8_t x_14; -x_14 = 1; -return x_14; +return x_22; } else { -size_t x_15; size_t x_16; -lean_dec(x_13); -x_15 = 1; -x_16 = lean_usize_add(x_5, x_15); -x_5 = x_16; -goto _start; +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_22, 0); +x_120 = lean_ctor_get(x_22, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_22); +x_121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +return x_121; } } } else { -uint8_t x_18; -x_18 = 0; -return x_18; +lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; +x_122 = lean_ctor_get(x_14, 1); +lean_inc(x_122); +lean_dec(x_14); +x_123 = lean_box(0); +x_124 = 0; +x_125 = lean_box(x_124); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_122); +lean_ctor_set(x_12, 1, x_126); +x_127 = lean_array_size(x_2); +x_128 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_4); +lean_inc(x_3); +x_129 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_123, x_2, x_127, x_128, x_12, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_130, 1); +lean_inc(x_131); +x_132 = lean_ctor_get(x_129, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_133 = x_129; +} else { + lean_dec_ref(x_129); + x_133 = lean_box(0); } +x_134 = lean_ctor_get(x_130, 0); +lean_inc(x_134); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_135 = x_130; +} else { + lean_dec_ref(x_130); + x_135 = lean_box(0); } +x_136 = lean_ctor_get(x_131, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_131, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_138 = x_131; +} else { + lean_dec_ref(x_131); + x_138 = lean_box(0); } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5) { -_start: -{ -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_4, x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_array_uget(x_3, x_4); -x_8 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); -if (lean_obj_tag(x_8) == 0) +x_139 = lean_unsigned_to_nat(0u); +x_140 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_139, x_134); +x_141 = lean_nat_dec_eq(x_140, x_1); +lean_dec(x_140); +if (x_141 == 0) { -size_t x_9; size_t x_10; -lean_dec(x_7); -x_9 = 1; -x_10 = lean_usize_add(x_4, x_9); -x_4 = x_10; -goto _start; -} -else +lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_138); +lean_dec(x_135); +lean_dec(x_133); +x_142 = lean_box(0); +x_143 = lean_unbox(x_136); +lean_dec(x_136); +lean_inc(x_5); +x_144 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_143, x_137, x_134, x_5, x_142, x_7, x_8, x_9, x_10, x_132); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +if (lean_obj_tag(x_145) == 0) { -lean_object* x_12; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -x_12 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_7); lean_dec(x_7); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = 1; -return x_13; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_146 = lean_ctor_get(x_144, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_147 = x_144; +} else { + lean_dec_ref(x_144); + x_147 = lean_box(0); +} +x_148 = lean_ctor_get(x_145, 0); +lean_inc(x_148); +lean_dec(x_145); +if (lean_is_scalar(x_147)) { + x_149 = lean_alloc_ctor(0, 2, 0); +} else { + x_149 = x_147; +} +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_146); +return x_149; } else { -size_t x_14; size_t x_15; -lean_dec(x_12); -x_14 = 1; -x_15 = lean_usize_add(x_4, x_14); -x_4 = x_15; +lean_object* x_150; lean_object* x_151; +x_150 = lean_ctor_get(x_144, 1); +lean_inc(x_150); +lean_dec(x_144); +x_151 = lean_ctor_get(x_145, 0); +lean_inc(x_151); +lean_dec(x_145); +x_6 = x_151; +x_11 = x_150; goto _start; } } -} else { -uint8_t x_17; -x_17 = 0; -return x_17; -} +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +if (lean_is_scalar(x_138)) { + x_153 = lean_alloc_ctor(0, 2, 0); +} else { + x_153 = x_138; } +lean_ctor_set(x_153, 0, x_136); +lean_ctor_set(x_153, 1, x_137); +if (lean_is_scalar(x_135)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_135; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_3 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_1); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_ctor_set(x_154, 0, x_134); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_154); +if (lean_is_scalar(x_133)) { + x_157 = lean_alloc_ctor(0, 2, 0); +} else { + x_157 = x_133; } +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_132); +return x_157; } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___closed__1; -x_5 = l_Lean_CollectFVars_main(x_3, x_4); -x_6 = lean_ctor_get(x_5, 2); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_array_get_size(x_6); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_lt(x_8, x_7); -if (x_9 == 0) -{ -uint8_t x_10; -lean_dec(x_7); -lean_dec(x_6); -x_10 = 1; -return x_10; } else { -size_t x_11; size_t x_12; uint8_t x_13; -x_11 = 0; -x_12 = lean_usize_of_nat(x_7); +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_13 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_6, x_11, x_12); -lean_dec(x_6); -if (x_13 == 0) -{ -uint8_t x_14; -x_14 = 1; -return x_14; -} -else -{ -uint8_t x_15; -x_15 = 0; -return x_15; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; uint8_t x_9; lean_object* x_10; -x_7 = lean_unbox_usize(x_5); lean_dec(x_5); -x_8 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_9 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1(x_1, x_2, x_3, x_4, x_7, x_8); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_10 = lean_box(x_9); -return x_10; -} +x_158 = lean_ctor_get(x_129, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_129, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_160 = x_129; +} else { + lean_dec_ref(x_129); + x_160 = lean_box(0); } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; uint8_t x_8; lean_object* x_9; -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_8 = l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__1___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___spec__2(x_1, x_2, x_3, x_6, x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_9 = lean_box(x_8); -return x_9; +if (lean_is_scalar(x_160)) { + x_161 = lean_alloc_ctor(1, 2, 0); +} else { + x_161 = x_160; } +lean_ctor_set(x_161, 0, x_158); +lean_ctor_set(x_161, 1, x_159); +return x_161; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -x_5 = lean_box(x_4); -return x_5; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1() { -_start: +else { -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 0; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; size_t x_171; size_t x_172; lean_object* x_173; +x_162 = lean_ctor_get(x_12, 1); +x_163 = lean_ctor_get(x_12, 0); +lean_inc(x_162); +lean_inc(x_163); +lean_dec(x_12); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +if (lean_is_exclusive(x_162)) { + lean_ctor_release(x_162, 0); + lean_ctor_release(x_162, 1); + x_165 = x_162; +} else { + lean_dec_ref(x_162); + x_165 = lean_box(0); } +x_166 = lean_box(0); +x_167 = 0; +x_168 = lean_box(x_167); +if (lean_is_scalar(x_165)) { + x_169 = lean_alloc_ctor(0, 2, 0); +} else { + x_169 = x_165; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2() { -_start: +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_164); +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_163); +lean_ctor_set(x_170, 1, x_169); +x_171 = lean_array_size(x_2); +x_172 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_4); +lean_inc(x_3); +x_173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_166, x_2, x_171, x_172, x_170, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_173) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; uint8_t x_185; +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +x_176 = lean_ctor_get(x_173, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_177 = x_173; +} else { + lean_dec_ref(x_173); + x_177 = lean_box(0); } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = lean_usize_dec_lt(x_8, x_7); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_14); -return x_16; +x_178 = lean_ctor_get(x_174, 0); +lean_inc(x_178); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_179 = x_174; +} else { + lean_dec_ref(x_174); + x_179 = lean_box(0); } -else -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_9); -x_17 = lean_array_uget(x_6, x_8); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_18 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +x_180 = lean_ctor_get(x_175, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_182 = x_175; +} else { + lean_dec_ref(x_175); + x_182 = lean_box(0); +} +x_183 = lean_unsigned_to_nat(0u); +x_184 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_183, x_178); +x_185 = lean_nat_dec_eq(x_184, x_1); +lean_dec(x_184); +if (x_185 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_18, 1); -x_22 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_20); -if (x_22 == 0) +lean_object* x_186; uint8_t x_187; lean_object* x_188; lean_object* x_189; +lean_dec(x_182); +lean_dec(x_179); +lean_dec(x_177); +x_186 = lean_box(0); +x_187 = lean_unbox(x_180); +lean_dec(x_180); +lean_inc(x_5); +x_188 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_187, x_181, x_178, x_5, x_186, x_7, x_8, x_9, x_10, x_176); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +if (lean_obj_tag(x_189) == 0) { -lean_object* x_23; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_5); -x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; -lean_ctor_set(x_18, 0, x_23); -return x_18; +lean_dec(x_4); +lean_dec(x_3); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +if (lean_is_exclusive(x_188)) { + lean_ctor_release(x_188, 0); + lean_ctor_release(x_188, 1); + x_191 = x_188; +} else { + lean_dec_ref(x_188); + x_191 = lean_box(0); +} +x_192 = lean_ctor_get(x_189, 0); +lean_inc(x_192); +lean_dec(x_189); +if (lean_is_scalar(x_191)) { + x_193 = lean_alloc_ctor(0, 2, 0); +} else { + x_193 = x_191; +} +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_190); +return x_193; } else { -size_t x_24; size_t x_25; -lean_free_object(x_18); -x_24 = 1; -x_25 = lean_usize_add(x_8, x_24); -lean_inc(x_5); -{ -size_t _tmp_7 = x_25; -lean_object* _tmp_8 = x_5; -lean_object* _tmp_13 = x_21; -x_8 = _tmp_7; -x_9 = _tmp_8; -x_14 = _tmp_13; -} +lean_object* x_194; lean_object* x_195; +x_194 = lean_ctor_get(x_188, 1); +lean_inc(x_194); +lean_dec(x_188); +x_195 = lean_ctor_get(x_189, 0); +lean_inc(x_195); +lean_dec(x_189); +x_6 = x_195; +x_11 = x_194; goto _start; } } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_18, 0); -x_28 = lean_ctor_get(x_18, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_18); -x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_27); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_5); -x_30 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__2; -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_28); -return x_31; +lean_dec(x_4); +lean_dec(x_3); +if (lean_is_scalar(x_182)) { + x_197 = lean_alloc_ctor(0, 2, 0); +} else { + x_197 = x_182; } -else -{ -size_t x_32; size_t x_33; -x_32 = 1; -x_33 = lean_usize_add(x_8, x_32); -lean_inc(x_5); -{ -size_t _tmp_7 = x_33; -lean_object* _tmp_8 = x_5; -lean_object* _tmp_13 = x_28; -x_8 = _tmp_7; -x_9 = _tmp_8; -x_14 = _tmp_13; +lean_ctor_set(x_197, 0, x_180); +lean_ctor_set(x_197, 1, x_181); +if (lean_is_scalar(x_179)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_179; } -goto _start; +lean_ctor_set(x_198, 0, x_178); +lean_ctor_set(x_198, 1, x_197); +x_199 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_198); +if (lean_is_scalar(x_177)) { + x_201 = lean_alloc_ctor(0, 2, 0); +} else { + x_201 = x_177; } +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_176); +return x_201; } } else { -uint8_t x_35; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_5); -x_35 = !lean_is_exclusive(x_18); -if (x_35 == 0) -{ -return x_18; +lean_dec(x_4); +lean_dec(x_3); +x_202 = lean_ctor_get(x_173, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_173, 1); +lean_inc(x_203); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_204 = x_173; +} else { + lean_dec_ref(x_173); + x_204 = lean_box(0); } -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_18, 0); -x_37 = lean_ctor_get(x_18, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_18); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +if (lean_is_scalar(x_204)) { + x_205 = lean_alloc_ctor(1, 2, 0); +} else { + x_205 = x_204; } +lean_ctor_set(x_205, 0, x_202); +lean_ctor_set(x_205, 1, x_203); +return x_205; } } } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("semiOutParam", 12, 12); -return x_1; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2() { -_start: +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_nat_dec_lt(x_7, x_15); +if (x_16 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_17; +lean_dec(x_7); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_14); +return x_17; } +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_array_fget(x_1, x_7); +x_19 = l_Lean_Expr_fvarId_x21(x_18); +lean_dec(x_18); +x_20 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_3, x_19); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_inc(x_7); +x_21 = lean_array_push(x_6, x_7); +x_22 = lean_ctor_get(x_5, 2); +x_23 = lean_nat_add(x_7, x_22); +lean_dec(x_7); +x_6 = x_21; +x_7 = x_23; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("outParam", 8, 8); -return x_1; +lean_object* x_25; lean_object* x_26; +lean_dec(x_20); +x_25 = lean_ctor_get(x_5, 2); +x_26 = lean_nat_add(x_7, x_25); +lean_dec(x_7); +x_7 = x_26; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -uint8_t x_15; -x_15 = lean_usize_dec_lt(x_8, x_7); +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_4, 1); +x_15 = lean_nat_dec_lt(x_6, x_14); if (x_15 == 0) { lean_object* x_16; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); +lean_dec(x_6); x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 0, x_5); +lean_ctor_set(x_16, 1, x_13); return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_27; -x_17 = lean_array_uget(x_6, x_8); -x_27 = !lean_is_exclusive(x_9); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_28 = lean_ctor_get(x_9, 1); -x_29 = lean_ctor_get(x_9, 0); -lean_dec(x_29); -x_30 = lean_ctor_get(x_28, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -x_32 = lean_ctor_get(x_28, 2); -lean_inc(x_32); -x_33 = lean_nat_dec_lt(x_31, x_32); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_array_fget(x_1, x_6); +x_18 = l_Lean_Expr_fvarId_x21(x_17); lean_dec(x_17); -lean_inc(x_4); -lean_ctor_set(x_9, 0, x_4); -x_34 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_34, 0, x_9); -x_18 = x_34; -x_19 = x_14; -goto block_26; +x_19 = l_Lean_RBNode_findCore___at_Lean_Meta_Closure_process___spec__1(x_2, x_18); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_inc(x_6); +x_20 = lean_array_push(x_5, x_6); +x_21 = lean_ctor_get(x_4, 2); +x_22 = lean_nat_add(x_6, x_21); +lean_dec(x_6); +x_5 = x_20; +x_6 = x_22; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; } else { -uint8_t x_35; -x_35 = !lean_is_exclusive(x_28); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_36 = lean_ctor_get(x_28, 2); -lean_dec(x_36); -x_37 = lean_ctor_get(x_28, 1); -lean_dec(x_37); -x_38 = lean_ctor_get(x_28, 0); -lean_dec(x_38); -x_39 = lean_array_fget(x_30, x_31); -x_40 = lean_unsigned_to_nat(1u); -x_41 = lean_nat_add(x_31, x_40); -lean_dec(x_31); -lean_ctor_set(x_28, 1, x_41); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_42 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_42) == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; -x_46 = l_Lean_Expr_isAppOf(x_43, x_45); -if (x_46 == 0) -{ -lean_object* x_47; uint8_t x_48; -x_47 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; -x_48 = l_Lean_Expr_isAppOf(x_43, x_47); -lean_dec(x_43); -if (x_48 == 0) +lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +x_24 = lean_ctor_get(x_4, 2); +x_25 = lean_nat_add(x_6, x_24); +lean_dec(x_6); +x_6 = x_25; +x_7 = lean_box(0); +x_8 = lean_box(0); +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -uint8_t x_49; -x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_39); -if (x_49 == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) { -lean_object* x_50; lean_object* x_51; -x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; -lean_ctor_set(x_9, 0, x_50); -x_51 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_51, 0, x_9); -x_18 = x_51; -x_19 = x_44; -goto block_26; +size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; +x_6 = 1; +x_7 = lean_usize_sub(x_2, x_6); +x_8 = lean_array_uget(x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +x_2 = x_7; +x_4 = x_9; +goto _start; } else { -lean_object* x_52; -lean_inc(x_4); -lean_ctor_set(x_9, 0, x_4); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_9); -x_18 = x_52; -x_19 = x_44; -goto block_26; +return x_4; +} +} } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_array_mk(x_1); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_2); +lean_ctor_set(x_14, 2, x_13); +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_3, x_4, x_14, x_14, x_11, x_12, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10); +lean_dec(x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_array_to_list(x_17); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_15, 0, x_19); +return x_15; } else { -lean_object* x_53; -lean_dec(x_39); -lean_inc(x_4); -lean_ctor_set(x_9, 0, x_4); -x_53 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_53, 0, x_9); -x_18 = x_53; -x_19 = x_44; -goto block_26; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_array_to_list(x_20); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +return x_24; } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_54; -lean_dec(x_43); -lean_dec(x_39); +lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = lean_box(0); +x_15 = 0; +x_16 = lean_box(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_14); +lean_ctor_set(x_19, 1, x_18); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); lean_inc(x_4); -lean_ctor_set(x_9, 0, x_4); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_9); -x_18 = x_54; -x_19 = x_44; -goto block_26; -} +x_20 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_2, x_3, x_4, x_5, x_14, x_19, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_box(0); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_4, x_6, x_3, x_25, x_26, x_9, x_10, x_11, x_12, x_24); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_25); +return x_27; } else { -uint8_t x_55; -lean_dec(x_28); -lean_dec(x_39); -lean_free_object(x_9); -lean_dec(x_13); +uint8_t x_28; +lean_dec(x_22); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); lean_dec(x_4); -x_55 = !lean_is_exclusive(x_42); -if (x_55 == 0) +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) { -return x_42; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_20, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_23, 0); +lean_inc(x_30); +lean_dec(x_23); +lean_ctor_set(x_20, 0, x_30); +return x_20; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_42, 0); -x_57 = lean_ctor_get(x_42, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_42); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_20, 1); +lean_inc(x_31); +lean_dec(x_20); +x_32 = lean_ctor_get(x_23, 0); +lean_inc(x_32); +lean_dec(x_23); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -lean_dec(x_28); -x_59 = lean_array_fget(x_30, x_31); -x_60 = lean_unsigned_to_nat(1u); -x_61 = lean_nat_add(x_31, x_60); -lean_dec(x_31); -x_62 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_62, 0, x_30); -lean_ctor_set(x_62, 1, x_61); -lean_ctor_set(x_62, 2, x_32); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_63 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -x_66 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; -x_67 = l_Lean_Expr_isAppOf(x_64, x_66); -if (x_67 == 0) -{ -lean_object* x_68; uint8_t x_69; -x_68 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; -x_69 = l_Lean_Expr_isAppOf(x_64, x_68); -lean_dec(x_64); -if (x_69 == 0) -{ -uint8_t x_70; -x_70 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_59); -if (x_70 == 0) +uint8_t x_34; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_4); +x_34 = !lean_is_exclusive(x_20); +if (x_34 == 0) { -lean_object* x_71; lean_object* x_72; -x_71 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; -lean_ctor_set(x_9, 1, x_62); -lean_ctor_set(x_9, 0, x_71); -x_72 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_72, 0, x_9); -x_18 = x_72; -x_19 = x_65; -goto block_26; +return x_20; } else { -lean_object* x_73; -lean_inc(x_4); -lean_ctor_set(x_9, 1, x_62); -lean_ctor_set(x_9, 0, x_4); -x_73 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_73, 0, x_9); -x_18 = x_73; -x_19 = x_65; -goto block_26; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_20, 0); +x_36 = lean_ctor_get(x_20, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_20); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } -else -{ -lean_object* x_74; -lean_dec(x_59); -lean_inc(x_4); -lean_ctor_set(x_9, 1, x_62); -lean_ctor_set(x_9, 0, x_4); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_9); -x_18 = x_74; -x_19 = x_65; -goto block_26; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1() { +_start: { -lean_object* x_75; -lean_dec(x_64); -lean_dec(x_59); -lean_inc(x_4); -lean_ctor_set(x_9, 1, x_62); -lean_ctor_set(x_9, 0, x_4); -x_75 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_75, 0, x_9); -x_18 = x_75; -x_19 = x_65; -goto block_26; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("numParams == xs.size\n ", 25, 25); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2() { +_start: { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -lean_dec(x_62); -lean_dec(x_59); -lean_free_object(x_9); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_76 = lean_ctor_get(x_63, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_63, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_78 = x_63; -} else { - lean_dec_ref(x_63); - x_78 = lean_box(0); -} -if (lean_is_scalar(x_78)) { - x_79 = lean_alloc_ctor(1, 2, 0); -} else { - x_79 = x_78; -} -lean_ctor_set(x_79, 0, x_76); -lean_ctor_set(x_79, 1, x_77); -return x_79; -} -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1; +x_3 = lean_string_append(x_1, x_2); +return x_3; } } -else -{ -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; -x_80 = lean_ctor_get(x_9, 1); -lean_inc(x_80); -lean_dec(x_9); -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -x_83 = lean_ctor_get(x_80, 2); -lean_inc(x_83); -x_84 = lean_nat_dec_lt(x_82, x_83); -if (x_84 == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3() { +_start: { -lean_object* x_85; lean_object* x_86; -lean_dec(x_83); -lean_dec(x_82); -lean_dec(x_81); -lean_dec(x_17); -lean_inc(x_4); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_4); -lean_ctor_set(x_85, 1, x_80); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -x_18 = x_86; -x_19 = x_14; -goto block_26; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.checkCoverage", 77, 77); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4() { +_start: { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - lean_ctor_release(x_80, 2); - x_87 = x_80; -} else { - lean_dec_ref(x_80); - x_87 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3; +x_3 = lean_unsigned_to_nat(357u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -x_88 = lean_array_fget(x_81, x_82); -x_89 = lean_unsigned_to_nat(1u); -x_90 = lean_nat_add(x_82, x_89); -lean_dec(x_82); -if (lean_is_scalar(x_87)) { - x_91 = lean_alloc_ctor(0, 3, 0); -} else { - x_91 = x_87; } -lean_ctor_set(x_91, 0, x_81); -lean_ctor_set(x_91, 1, x_90); -lean_ctor_set(x_91, 2, x_83); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_92 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_92) == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; -x_93 = lean_ctor_get(x_92, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_92, 1); -lean_inc(x_94); -lean_dec(x_92); -x_95 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__2; -x_96 = l_Lean_Expr_isAppOf(x_93, x_95); -if (x_96 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_array_get_size(x_3); +x_11 = lean_nat_dec_eq(x_1, x_10); +if (x_11 == 0) { -lean_object* x_97; uint8_t x_98; -x_97 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___closed__4; -x_98 = l_Lean_Expr_isAppOf(x_93, x_97); -lean_dec(x_93); -if (x_98 == 0) +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_3); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4; +x_13 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(x_12, x_5, x_6, x_7, x_8, x_9); +return x_13; +} +else { -uint8_t x_99; -x_99 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_1, x_2, x_88); -if (x_99 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_48; uint8_t x_49; +x_14 = lean_box(0); +x_15 = lean_ctor_get(x_2, 1); +x_16 = lean_array_get_size(x_15); +lean_inc(x_3); +x_17 = lean_array_to_list(x_3); +x_18 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(x_17, x_14); +x_19 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_18); +x_20 = lean_box(0); +x_48 = lean_unsigned_to_nat(0u); +x_49 = lean_nat_dec_lt(x_48, x_16); +if (x_49 == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___closed__1; -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_101, 1, x_91); -x_102 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_102, 0, x_101); -x_18 = x_102; -x_19 = x_94; -goto block_26; +lean_dec(x_16); +x_21 = x_14; +goto block_47; } else { -lean_object* x_103; lean_object* x_104; -lean_inc(x_4); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_4); -lean_ctor_set(x_103, 1, x_91); -x_104 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_104, 0, x_103); -x_18 = x_104; -x_19 = x_94; -goto block_26; +size_t x_50; size_t x_51; lean_object* x_52; +x_50 = lean_usize_of_nat(x_16); +lean_dec(x_16); +x_51 = 0; +x_52 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_15, x_50, x_51, x_14); +x_21 = x_52; +goto block_47; } +block_47: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_3, x_10, x_21, x_14); +lean_inc(x_22); +x_23 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_22); +lean_inc(x_5); +lean_inc(x_23); +lean_inc(x_22); +x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_14, x_19, x_20, x_22, x_22, x_22, x_23, lean_box(0), x_5, x_6, x_7, x_8, x_9); +lean_dec(x_22); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_28, x_26); +x_30 = lean_nat_dec_eq(x_29, x_1); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_free_object(x_24); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_26, x_31, x_5, x_6, x_7, x_8, x_27); +lean_dec(x_3); +return x_32; } else { -lean_object* x_105; lean_object* x_106; -lean_dec(x_88); -lean_inc(x_4); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_4); -lean_ctor_set(x_105, 1, x_91); -x_106 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_106, 0, x_105); -x_18 = x_106; -x_19 = x_94; -goto block_26; +lean_object* x_33; +lean_dec(x_26); +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_33 = lean_box(0); +lean_ctor_set(x_24, 0, x_33); +return x_24; } } else { -lean_object* x_107; lean_object* x_108; -lean_dec(x_93); -lean_dec(x_88); -lean_inc(x_4); -x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_4); -lean_ctor_set(x_107, 1, x_91); -x_108 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_108, 0, x_107); -x_18 = x_108; -x_19 = x_94; -goto block_26; -} +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_34 = lean_ctor_get(x_24, 0); +x_35 = lean_ctor_get(x_24, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_24); +x_36 = lean_unsigned_to_nat(0u); +x_37 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_36, x_34); +x_38 = lean_nat_dec_eq(x_37, x_1); +lean_dec(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_box(0); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_34, x_39, x_5, x_6, x_7, x_8, x_35); +lean_dec(x_3); +return x_40; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_91); -lean_dec(x_88); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_41; lean_object* x_42; +lean_dec(x_34); +lean_dec(x_23); +lean_dec(x_19); lean_dec(x_10); -lean_dec(x_4); -x_109 = lean_ctor_get(x_92, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_92, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_111 = x_92; -} else { - lean_dec_ref(x_92); - x_111 = lean_box(0); -} -if (lean_is_scalar(x_111)) { - x_112 = lean_alloc_ctor(1, 2, 0); -} else { - x_112 = x_111; -} -lean_ctor_set(x_112, 0, x_109); -lean_ctor_set(x_112, 1, x_110); -return x_112; -} +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_35); +return x_42; } } -block_26: -{ -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; } else { -lean_object* x_22; size_t x_23; size_t x_24; -x_22 = lean_ctor_get(x_18, 0); -lean_inc(x_22); -lean_dec(x_18); -x_23 = 1; -x_24 = lean_usize_add(x_8, x_23); -x_8 = x_24; -x_9 = x_22; -x_14 = x_19; -goto _start; +uint8_t x_43; +lean_dec(x_23); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_43 = !lean_is_exclusive(x_24); +if (x_43 == 0) +{ +return x_24; } +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_24, 0); +x_45 = lean_ctor_get(x_24, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_24); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; lean_object* x_8; lean_object* x_9; -x_7 = 1; -x_8 = lean_box(x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_6); -return x_9; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed), 6, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +lean_object* x_10; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_10 = lean_infer_type(x_1, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; -x_14 = lean_array_get_size(x_1); -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Array_toSubarray___rarg(x_1, x_15, x_14); -lean_inc(x_2); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_2); -lean_ctor_set(x_17, 1, x_16); -x_18 = lean_array_size(x_7); -lean_inc(x_12); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_3, x_4, x_5, x_2, x_7, x_7, x_18, x_6, x_17, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -lean_dec(x_20); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___closed__1; -x_24 = lean_box(0); -x_25 = lean_apply_6(x_23, x_24, x_9, x_10, x_11, x_12, x_22); -return x_25; +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +lean_inc(x_2); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_2); +x_14 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed), 9, 2); +lean_closure_set(x_14, 0, x_2); +lean_closure_set(x_14, 1, x_3); +x_15 = 0; +x_16 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_11, x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_12); +return x_16; } else { -uint8_t x_26; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) +uint8_t x_17; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_10); +if (x_17 == 0) { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_19, 0); -lean_dec(x_27); -x_28 = lean_ctor_get(x_21, 0); -lean_inc(x_28); -lean_dec(x_21); -lean_ctor_set(x_19, 0, x_28); -return x_19; +return x_10; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_19, 1); -lean_inc(x_29); -lean_dec(x_19); -x_30 = lean_ctor_get(x_21, 0); -lean_inc(x_30); -lean_dec(x_21); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -return x_31; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_10, 0); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_10); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -uint8_t x_32; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +x_10 = lean_nat_dec_eq(x_9, x_2); lean_dec(x_9); -x_32 = !lean_is_exclusive(x_19); -if (x_32 == 0) +if (x_10 == 0) { -return x_19; +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_11, x_4, x_5, x_6, x_7, x_8); +return x_12; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_19, 0); -x_34 = lean_ctor_get(x_19, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_19); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_13; lean_object* x_14; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_8); +return x_14; } } } +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_15; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_15 = lean_infer_type(x_1, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_15) == 0) +lean_object* x_6; +x_6 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unsigned_to_nat(0u); -x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_18); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = lean_box_usize(x_8); -x_22 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed), 13, 6); -lean_closure_set(x_22, 0, x_3); -lean_closure_set(x_22, 1, x_4); -lean_closure_set(x_22, 2, x_5); -lean_closure_set(x_22, 3, x_6); -lean_closure_set(x_22, 4, x_7); -lean_closure_set(x_22, 5, x_21); -x_23 = 0; -x_24 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_16, x_20, x_22, x_23, x_10, x_11, x_12, x_13, x_17); -return x_24; +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -uint8_t x_25; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) -{ +lean_dec(x_2); return x_15; } -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; -} -} } -} -static lean_object* _init_l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1() { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_14; +x_14 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -if (lean_obj_tag(x_5) == 5) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_13 = lean_ctor_get(x_5, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_5, 1); -lean_inc(x_14); +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); lean_dec(x_5); -x_15 = lean_array_set(x_6, x_7, x_14); -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_7, x_16); -lean_dec(x_7); -x_5 = x_13; -x_6 = x_15; -x_7 = x_17; -goto _start; +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } -else -{ -lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_7); -x_19 = lean_box(0); -x_20 = lean_box(0); -x_21 = lean_array_size(x_3); -x_22 = 0; -x_23 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_24 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_19, x_23, x_3, x_21, x_22, x_23, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -lean_dec(x_25); -if (lean_obj_tag(x_26) == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = lean_box(0); -x_29 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_5, x_4, x_6, x_20, x_1, x_2, x_19, x_22, x_28, x_8, x_9, x_10, x_11, x_27); -return x_29; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -uint8_t x_30; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_30 = !lean_is_exclusive(x_24); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_24, 0); -lean_dec(x_31); -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -lean_ctor_set(x_24, 0, x_32); -return x_24; +return x_9; } -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_24, 1); -lean_inc(x_33); -lean_dec(x_24); -x_34 = lean_ctor_get(x_26, 0); -lean_inc(x_34); -lean_dec(x_26); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -return x_35; } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; } } -else +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -uint8_t x_36; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = lean_unbox_usize(x_8); lean_dec(x_8); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_24); -if (x_36 == 0) -{ -return x_24; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_24, 0); -x_38 = lean_ctor_get(x_24, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_24); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; -} -} -} +return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_10); -x_12 = l_Lean_Meta_Grind_ppPattern___closed__5; -lean_inc(x_11); -x_13 = lean_mk_array(x_11, x_12); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_sub(x_11, x_14); -lean_dec(x_11); -lean_inc(x_4); -x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_4, x_13, x_15, x_5, x_6, x_7, x_8, x_9); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_1, x_2, x_3, x_4, x_5, x_14, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_5); lean_dec(x_4); +lean_dec(x_1); return x_16; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_9; uint8_t x_10; lean_object* x_11; -x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1___boxed), 9, 2); -lean_closure_set(x_9, 0, x_1); -lean_closure_set(x_9, 1, x_2); -x_10 = 0; -x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_3, x_9, x_10, x_4, x_5, x_6, x_7, x_8); -return x_11; +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_1); +lean_dec(x_1); +x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); +lean_object* x_13; +x_13 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_17; +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +lean_dec(x_1); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); +lean_object* x_15; +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_17; +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_7; -x_7 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); +lean_object* x_14; +x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_1, x_5, x_6, x_4); +lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -size_t x_14; lean_object* x_15; -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__2(x_1, x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_15; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -size_t x_15; lean_object* x_16; -x_15 = lean_unbox_usize(x_8); +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); -x_16 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_15, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_9); +lean_dec(x_3); lean_dec(x_2); -return x_16; +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_13; -x_13 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_4); -lean_dec(x_3); -return x_13; +lean_dec(x_2); +lean_dec(x_1); +return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_3); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); return x_10; } } -static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instInhabitedMetaM___boxed), 5, 1); -lean_closure_set(x_1, 0, lean_box(0)); +x_1 = lean_mk_string_unchecked(" : ", 3, 3); return x_1; } } -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2() { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; -x_8 = lean_panic_fn(x_7, x_1); -x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); -return x_9; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -if (lean_obj_tag(x_2) == 0) -{ +lean_object* x_11; lean_inc(x_1); -return x_1; +x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_13 = lean_ctor_get(x_11, 0); +x_14 = l_Lean_MessageData_ofExpr(x_1); +lean_inc(x_2); +x_15 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_15, 0, x_2); +lean_ctor_set(x_15, 1, x_14); +x_16 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean_MessageData_ofExpr(x_13); +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_2); +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_4); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_box(x_3); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_11, 0, x_24); +return x_11; } else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_2, 0); -x_4 = lean_ctor_get(x_2, 2); -x_5 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_4); -lean_inc(x_3); -x_6 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_6, 0, x_3); -lean_ctor_set(x_6, 1, x_5); -return x_6; -} +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_25 = lean_ctor_get(x_11, 0); +x_26 = lean_ctor_get(x_11, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_11); +x_27 = l_Lean_MessageData_ofExpr(x_1); +lean_inc(x_2); +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_2); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_MessageData_ofExpr(x_25); +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_2); +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_4); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_box(x_3); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_26); +return x_38; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_1) == 0) +uint8_t x_39; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_11); +if (x_39 == 0) { -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; +return x_11; } else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = l_Lean_Expr_fvarId_x21(x_5); -lean_dec(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_11, 0); +x_41 = lean_ctor_get(x_11, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_11); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1() { +_start: { -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n", 1, 1); +return x_1; } -goto _start; } -else +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2() { +_start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = l_Lean_Expr_fvarId_x21(x_9); -lean_dec(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2; +x_2 = l_Lean_MessageData_ofFormat(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_nat_dec_lt(x_7, x_15); +if (x_16 == 0) { -lean_object* x_2; -x_2 = lean_box(0); -return x_2; +lean_object* x_17; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_14); +return x_17; } else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_4); -x_6 = lean_box(0); -x_7 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_3, x_6); -return x_7; -} -} -} -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +uint8_t x_18; +x_18 = !lean_is_exclusive(x_6); +if (x_18 == 0) { -if (lean_obj_tag(x_4) == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_6, 0); +x_20 = lean_ctor_get(x_6, 1); +x_21 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); +if (x_21 == 0) { -lean_object* x_6; -x_6 = l_List_reverse___rarg(x_5); -return x_6; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_5, 2); +x_23 = lean_nat_add(x_7, x_22); +lean_dec(x_7); +x_7 = x_23; +x_8 = lean_box(0); +x_9 = lean_box(0); +goto _start; } else { -uint8_t x_7; -x_7 = !lean_is_exclusive(x_4); -if (x_7 == 0) +lean_object* x_25; uint8_t x_26; +lean_free_object(x_6); +x_25 = lean_array_fget(x_2, x_7); +x_26 = lean_unbox(x_19); +if (x_26 == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_8 = lean_ctor_get(x_4, 0); -x_9 = lean_ctor_get(x_4, 1); -x_10 = lean_nat_sub(x_1, x_8); -lean_dec(x_8); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_sub(x_10, x_11); -lean_dec(x_10); -x_13 = lean_nat_dec_lt(x_12, x_3); -if (x_13 == 0) +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_27 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_20); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_box(0); +x_30 = lean_unbox(x_19); +lean_dec(x_19); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_31 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_30, x_28, x_29, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_32; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 0) +{ +uint8_t x_33; +lean_dec(x_13); lean_dec(x_12); -x_14 = l_Lean_instInhabitedExpr; -x_15 = l_outOfBounds___rarg(x_14); -x_16 = l_Lean_Expr_fvarId_x21(x_15); -lean_dec(x_15); -lean_ctor_set(x_4, 1, x_5); -lean_ctor_set(x_4, 0, x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_33 = !lean_is_exclusive(x_31); +if (x_33 == 0) { -lean_object* _tmp_3 = x_9; -lean_object* _tmp_4 = x_4; -x_4 = _tmp_3; -x_5 = _tmp_4; +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_31, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +lean_ctor_set(x_31, 0, x_35); +return x_31; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); +lean_dec(x_31); +x_37 = lean_ctor_get(x_32, 0); +lean_inc(x_37); +lean_dec(x_32); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; } +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_31, 1); +lean_inc(x_39); +lean_dec(x_31); +x_40 = lean_ctor_get(x_32, 0); +lean_inc(x_40); +lean_dec(x_32); +x_41 = lean_ctor_get(x_5, 2); +x_42 = lean_nat_add(x_7, x_41); +lean_dec(x_7); +x_6 = x_40; +x_7 = x_42; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_14 = x_39; goto _start; } +} else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_array_fget(x_2, x_12); +uint8_t x_44; +lean_dec(x_13); lean_dec(x_12); -x_19 = l_Lean_Expr_fvarId_x21(x_18); -lean_dec(x_18); -lean_ctor_set(x_4, 1, x_5); -lean_ctor_set(x_4, 0, x_19); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_44 = !lean_is_exclusive(x_31); +if (x_44 == 0) { -lean_object* _tmp_3 = x_9; -lean_object* _tmp_4 = x_4; -x_4 = _tmp_3; -x_5 = _tmp_4; +return x_31; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_31, 0); +x_46 = lean_ctor_get(x_31, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_31); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } -goto _start; } } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_21 = lean_ctor_get(x_4, 0); -x_22 = lean_ctor_get(x_4, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_4); -x_23 = lean_nat_sub(x_1, x_21); -lean_dec(x_21); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_sub(x_23, x_24); -lean_dec(x_23); -x_26 = lean_nat_dec_lt(x_25, x_3); -if (x_26 == 0) +uint8_t x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_19); +x_48 = 0; +x_49 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_50 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_48, x_20, x_49, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_25); -x_27 = l_Lean_instInhabitedExpr; -x_28 = l_outOfBounds___rarg(x_27); -x_29 = l_Lean_Expr_fvarId_x21(x_28); -lean_dec(x_28); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_5); -x_4 = x_22; -x_5 = x_30; -goto _start; +lean_object* x_51; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +if (lean_obj_tag(x_51) == 0) +{ +uint8_t x_52; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_52 = !lean_is_exclusive(x_50); +if (x_52 == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_50, 0); +lean_dec(x_53); +x_54 = lean_ctor_get(x_51, 0); +lean_inc(x_54); +lean_dec(x_51); +lean_ctor_set(x_50, 0, x_54); +return x_50; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_array_fget(x_2, x_25); -lean_dec(x_25); -x_33 = l_Lean_Expr_fvarId_x21(x_32); -lean_dec(x_32); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_5); -x_4 = x_22; -x_5 = x_34; -goto _start; -} -} -} +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_50, 1); +lean_inc(x_55); +lean_dec(x_50); +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +lean_dec(x_51); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +return x_57; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) -{ -lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_8 = lean_array_uget(x_3, x_4); -x_9 = 1; -x_10 = lean_usize_add(x_4, x_9); -x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); -if (lean_obj_tag(x_11) == 0) +else { -lean_dec(x_8); -x_4 = x_10; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_50, 1); +lean_inc(x_58); +lean_dec(x_50); +x_59 = lean_ctor_get(x_51, 0); +lean_inc(x_59); +lean_dec(x_51); +x_60 = lean_ctor_get(x_5, 2); +x_61 = lean_nat_add(x_7, x_60); +lean_dec(x_7); +x_6 = x_59; +x_7 = x_61; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_14 = x_58; goto _start; } +} else { -lean_object* x_13; lean_object* x_14; +uint8_t x_63; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); -x_13 = lean_box(0); -x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); -x_4 = x_10; -x_6 = x_14; -goto _start; -} +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_63 = !lean_is_exclusive(x_50); +if (x_63 == 0) +{ +return x_50; } else { -return x_6; +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_50, 0); +x_65 = lean_ctor_get(x_50, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_50); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { -_start: -{ -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) +} +} +else { -lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); -if (lean_obj_tag(x_10) == 0) +lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = lean_ctor_get(x_6, 0); +x_68 = lean_ctor_get(x_6, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_6); +x_69 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); +if (x_69 == 0) { +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_67); +lean_ctor_set(x_70, 1, x_68); +x_71 = lean_ctor_get(x_5, 2); +x_72 = lean_nat_add(x_7, x_71); lean_dec(x_7); -x_3 = x_9; +x_6 = x_70; +x_7 = x_72; +x_8 = lean_box(0); +x_9 = lean_box(0); goto _start; } else { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -x_12 = lean_box(0); -x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); -x_3 = x_9; -x_5 = x_13; -goto _start; -} -} -else +lean_object* x_74; uint8_t x_75; +x_74 = lean_array_fget(x_2, x_7); +x_75 = lean_unbox(x_67); +if (x_75 == 0) { -return x_5; -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; +x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_68); +lean_ctor_set(x_77, 1, x_76); +x_78 = lean_box(0); +x_79 = lean_unbox(x_67); +lean_dec(x_67); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_80 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_79, x_77, x_78, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_80) == 0) { -if (lean_obj_tag(x_7) == 0) +lean_object* x_81; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (lean_obj_tag(x_81) == 0) { -lean_object* x_15; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_14); -return x_15; +lean_dec(x_7); +lean_dec(x_3); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_83 = x_80; +} else { + lean_dec_ref(x_80); + x_83 = lean_box(0); +} +x_84 = lean_ctor_get(x_81, 0); +lean_inc(x_84); +lean_dec(x_81); +if (lean_is_scalar(x_83)) { + x_85 = lean_alloc_ctor(0, 2, 0); +} else { + x_85 = x_83; +} +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +return x_85; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_7, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_7, 1); -lean_inc(x_17); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_86 = lean_ctor_get(x_80, 1); +lean_inc(x_86); +lean_dec(x_80); +x_87 = lean_ctor_get(x_81, 0); +lean_inc(x_87); +lean_dec(x_81); +x_88 = lean_ctor_get(x_5, 2); +x_89 = lean_nat_add(x_7, x_88); lean_dec(x_7); -lean_inc(x_10); -x_18 = l_Lean_FVarId_getType(x_16, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_box(0); -lean_inc(x_1); -x_22 = lean_array_mk(x_1); -x_23 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_24 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_22); -x_25 = l_Lean_CollectFVars_main(x_19, x_24); -x_26 = lean_ctor_get(x_25, 2); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_array_get_size(x_26); -x_28 = lean_unsigned_to_nat(0u); -x_29 = lean_nat_dec_lt(x_28, x_27); -if (x_29 == 0) -{ -lean_dec(x_27); -lean_dec(x_26); -x_7 = x_17; +x_6 = x_87; +x_7 = x_89; +x_8 = lean_box(0); x_9 = lean_box(0); -x_14 = x_20; +x_14 = x_86; goto _start; } +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_91 = lean_ctor_get(x_80, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_80, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_93 = x_80; +} else { + lean_dec_ref(x_80); + x_93 = lean_box(0); +} +if (lean_is_scalar(x_93)) { + x_94 = lean_alloc_ctor(1, 2, 0); +} else { + x_94 = x_93; +} +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_92); +return x_94; +} +} else { -uint8_t x_31; -x_31 = lean_nat_dec_le(x_27, x_27); -if (x_31 == 0) +uint8_t x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_67); +x_95 = 0; +x_96 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_3); +x_97 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_95, x_68, x_96, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_97) == 0) { -lean_dec(x_27); -lean_dec(x_26); -x_7 = x_17; -x_9 = lean_box(0); -x_14 = x_20; -goto _start; +lean_object* x_98; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_3); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_100 = x_97; +} else { + lean_dec_ref(x_97); + x_100 = lean_box(0); +} +x_101 = lean_ctor_get(x_98, 0); +lean_inc(x_101); +lean_dec(x_98); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; +} +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } else { -size_t x_33; size_t x_34; lean_object* x_35; -x_33 = 0; -x_34 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_35 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_3, x_26, x_33, x_34, x_8); -lean_dec(x_26); -x_7 = x_17; -x_8 = x_35; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_103 = lean_ctor_get(x_97, 1); +lean_inc(x_103); +lean_dec(x_97); +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +lean_dec(x_98); +x_105 = lean_ctor_get(x_5, 2); +x_106 = lean_nat_add(x_7, x_105); +lean_dec(x_7); +x_6 = x_104; +x_7 = x_106; +x_8 = lean_box(0); x_9 = lean_box(0); -x_14 = x_20; +x_14 = x_103; goto _start; } } -} else { -uint8_t x_37; -lean_dec(x_17); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_1); -x_37 = !lean_is_exclusive(x_18); -if (x_37 == 0) -{ -return x_18; +lean_dec(x_7); +lean_dec(x_3); +x_108 = lean_ctor_get(x_97, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_97, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_110 = x_97; +} else { + lean_dec_ref(x_97); + x_110 = lean_box(0); } -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_18, 0); -x_39 = lean_ctor_get(x_18, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_18); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; } +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; } } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_14; -lean_dec(x_9); -lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_7); -lean_ctor_set(x_14, 1, x_13); -return x_14; } -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_6, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_6, 1); -lean_inc(x_16); -lean_dec(x_6); -lean_inc(x_9); -x_17 = l_Lean_FVarId_getType(x_15, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_box(0); -lean_inc(x_1); -x_21 = lean_array_mk(x_1); -x_22 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_23 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_20); -lean_ctor_set(x_23, 2, x_21); -x_24 = l_Lean_CollectFVars_main(x_18, x_23); -x_25 = lean_ctor_get(x_24, 2); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_array_get_size(x_25); -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_nat_dec_lt(x_27, x_26); -if (x_28 == 0) -{ -lean_dec(x_26); -lean_dec(x_25); -x_6 = x_16; -x_8 = lean_box(0); -x_13 = x_19; -goto _start; } -else -{ -uint8_t x_30; -x_30 = lean_nat_dec_le(x_26, x_26); -if (x_30 == 0) -{ -lean_dec(x_26); -lean_dec(x_25); -x_6 = x_16; -x_8 = lean_box(0); -x_13 = x_19; -goto _start; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1() { +_start: { -size_t x_32; size_t x_33; lean_object* x_34; -x_32 = 0; -x_33 = lean_usize_of_nat(x_26); -lean_dec(x_26); -x_34 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_2, x_25, x_32, x_33, x_7); -lean_dec(x_25); -x_6 = x_16; -x_7 = x_34; -x_8 = lean_box(0); -x_13 = x_19; -goto _start; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_array_get_size(x_2); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_9); +lean_ctor_set(x_12, 2, x_11); +x_13 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_13, x_12, x_12, x_14, x_10, lean_box(0), lean_box(0), x_4, x_5, x_6, x_7, x_8); +lean_dec(x_12); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_18, x_4, x_5, x_6, x_7, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_19; } else { -uint8_t x_36; -lean_dec(x_16); -lean_dec(x_9); +uint8_t x_20; lean_dec(x_7); -lean_dec(x_1); -x_36 = !lean_is_exclusive(x_17); -if (x_36 == 0) +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) { -return x_17; +return x_15; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_17, 0); -x_38 = lean_ctor_get(x_17, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_17); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) -{ -lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_8 = lean_array_uget(x_3, x_4); -x_9 = 1; -x_10 = lean_usize_add(x_4, x_9); -x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); -if (lean_obj_tag(x_11) == 0) +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = lean_infer_type(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) { -lean_dec(x_8); -x_4 = x_10; -goto _start; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_2); +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed), 8, 1); +lean_closure_set(x_13, 0, x_3); +x_14 = 0; +x_15 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_10, x_12, x_13, x_14, x_4, x_5, x_6, x_7, x_11); +return x_15; } else { -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); -x_13 = lean_box(0); -x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); -x_4 = x_10; -x_6 = x_14; -goto _start; -} +uint8_t x_16; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_16 = !lean_is_exclusive(x_9); +if (x_16 == 0) +{ +return x_9; } else { -return x_6; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_9, 0); +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_9); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); -if (lean_obj_tag(x_10) == 0) -{ -lean_dec(x_7); -x_3 = x_9; -goto _start; +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -x_12 = lean_box(0); -x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); -x_3 = x_9; -x_5 = x_13; -goto _start; +lean_object* x_15; +x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_15; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -return x_5; -} +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) -{ -lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_8 = lean_array_uget(x_3, x_4); -x_9 = 1; -x_10 = lean_usize_add(x_4, x_9); -x_11 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_8); -if (lean_obj_tag(x_11) == 0) +if (lean_obj_tag(x_1) == 0) { -lean_dec(x_8); -x_4 = x_10; -goto _start; +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; } else { -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); -x_13 = lean_box(0); -x_14 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_6, x_8, x_13); -x_4 = x_10; -x_6 = x_14; -goto _start; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Lean_Meta_Grind_ppPattern(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; } +goto _start; } else { -return x_6; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_Meta_Grind_ppPattern(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_10 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_1, x_7); -if (lean_obj_tag(x_10) == 0) +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else { -lean_dec(x_7); -x_3 = x_9; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_1, 1); +lean_ctor_set(x_1, 1, x_2); +{ +lean_object* _tmp_0 = x_5; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} goto _start; } else { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -x_12 = lean_box(0); -x_13 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_5, x_7, x_12); -x_3 = x_9; -x_5 = x_13; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_9 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_2); +x_1 = x_8; +x_2 = x_9; goto _start; } } -else -{ -return x_5; -} } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_15; -x_15 = lean_usize_dec_lt(x_8, x_7); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_14); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_25; -x_17 = lean_array_uget(x_6, x_8); -x_25 = !lean_is_exclusive(x_9); -if (x_25 == 0) -{ -lean_object* x_26; uint8_t x_27; -x_26 = lean_ctor_get(x_9, 1); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +switch (lean_obj_tag(x_1)) { +case 0: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_28 = lean_ctor_get(x_9, 0); -x_29 = lean_ctor_get(x_26, 0); -x_30 = lean_ctor_get(x_26, 1); -x_31 = l_Lean_Expr_fvarId_x21(x_17); -x_32 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_30, x_31); -if (lean_obj_tag(x_32) == 0) +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +x_8 = l_Lean_mkConstWithLevelParams___at_Lean_Meta_registerCoercion___spec__1(x_7, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_33; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_33 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_33) == 0) +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_28, x_31); -if (lean_obj_tag(x_36) == 0) +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +x_11 = l_Lean_MessageData_ofConst(x_10); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else { -lean_object* x_37; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_34); -x_37 = l_Lean_Meta_isProp(x_34, x_10, x_11, x_12, x_13, x_35); -if (lean_obj_tag(x_37) == 0) +lean_dec(x_8); +x_14 = l_Lean_MessageData_ofConst(x_12); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +} +else { -lean_object* x_38; uint8_t x_39; -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_unbox(x_38); -lean_dec(x_38); -if (x_39 == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_8); +if (x_16 == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_37, 1); -lean_inc(x_40); -lean_dec(x_37); -lean_inc(x_10); -lean_inc(x_31); -x_41 = l_Lean_FVarId_getDecl(x_31, x_10, x_11, x_12, x_13, x_40); -if (lean_obj_tag(x_41) == 0) +return x_8; +} +else { -lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_LocalDecl_binderInfo(x_42); -lean_dec(x_42); -x_45 = lean_box(x_44); -if (lean_obj_tag(x_45) == 3) +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_8, 0); +x_18 = lean_ctor_get(x_8, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_8); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +case 1: { -lean_object* x_46; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_34); -lean_inc(x_28); -lean_inc(x_4); -x_46 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_28, x_34, x_10, x_11, x_12, x_13, x_43); -if (lean_obj_tag(x_46) == 0) +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = l_Lean_Expr_fvar___override(x_20); +x_22 = l_Lean_MessageData_ofExpr(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_6); +return x_23; +} +case 2: { -lean_object* x_47; uint8_t x_48; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_unbox(x_47); -lean_dec(x_47); -if (x_48 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_1); +if (x_24 == 0) { -lean_object* x_49; lean_object* x_50; -lean_dec(x_34); -lean_dec(x_31); -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); -x_50 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_50, 0, x_9); -x_18 = x_50; -x_19 = x_49; -goto block_24; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_1, 1); +x_26 = lean_ctor_get(x_1, 0); +lean_dec(x_26); +x_27 = l_Lean_MessageData_ofSyntax(x_25); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 1, x_6); +lean_ctor_set(x_1, 0, x_27); +return x_1; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; -lean_dec(x_29); -x_51 = lean_ctor_get(x_46, 1); -lean_inc(x_51); -lean_dec(x_46); -x_52 = lean_box(0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_1, 1); +lean_inc(x_28); +lean_dec(x_1); +x_29 = l_Lean_MessageData_ofSyntax(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_6); +return x_30; +} +} +default: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_1, 0); lean_inc(x_31); -x_53 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_52); -x_54 = lean_box(0); -lean_inc(x_2); -x_55 = lean_array_mk(x_2); -x_56 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_57 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -lean_ctor_set(x_57, 2, x_55); -x_58 = l_Lean_CollectFVars_main(x_34, x_57); -x_59 = lean_ctor_get(x_58, 2); -lean_inc(x_59); -lean_dec(x_58); -x_60 = lean_array_get_size(x_59); -x_61 = lean_unsigned_to_nat(0u); -x_62 = lean_nat_dec_lt(x_61, x_60); -x_63 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_52); -if (x_62 == 0) +lean_dec(x_1); +x_32 = l_Lean_MessageData_ofName(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_6); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_64; lean_object* x_65; lean_object* x_66; -lean_dec(x_60); -lean_dec(x_59); -x_64 = 1; -x_65 = lean_box(x_64); -lean_ctor_set(x_26, 1, x_63); -lean_ctor_set(x_26, 0, x_65); -lean_ctor_set(x_9, 0, x_53); -x_66 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_66, 0, x_9); -x_18 = x_66; -x_19 = x_51; -goto block_24; +lean_object* x_13; lean_object* x_14; +x_13 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +lean_ctor_set(x_13, 2, x_3); +lean_ctor_set(x_13, 3, x_4); +lean_ctor_set(x_13, 4, x_5); +lean_ctor_set(x_13, 5, x_6); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1() { +_start: { -uint8_t x_67; -x_67 = lean_nat_dec_le(x_60, x_60); -if (x_67 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid pattern(s) for `", 24, 24); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2() { +_start: { -uint8_t x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_60); -lean_dec(x_59); -x_68 = 1; -x_69 = lean_box(x_68); -lean_ctor_set(x_26, 1, x_63); -lean_ctor_set(x_26, 0, x_69); -lean_ctor_set(x_9, 0, x_53); -x_70 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_70, 0, x_9); -x_18 = x_70; -x_19 = x_51; -goto block_24; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5() { +_start: { -size_t x_71; size_t x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; -x_71 = 0; -x_72 = lean_usize_of_nat(x_60); -lean_dec(x_60); -x_73 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_59, x_71, x_72, x_53); -lean_dec(x_59); -x_74 = 1; -x_75 = lean_box(x_74); -lean_ctor_set(x_26, 1, x_63); -lean_ctor_set(x_26, 0, x_75); -lean_ctor_set(x_9, 0, x_73); -x_76 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_76, 0, x_9); -x_18 = x_76; -x_19 = x_51; -goto block_24; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\nthe following theorem parameters cannot be instantiated:", 57, 57); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_77; -lean_dec(x_34); -lean_dec(x_31); -lean_free_object(x_26); -lean_dec(x_30); -lean_dec(x_29); -lean_free_object(x_9); -lean_dec(x_28); -lean_dec(x_13); +lean_object* x_14; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(x_1, x_2, x_3, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_box(0); +x_18 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1(x_4, x_1, x_2, x_5, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_16); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_77 = !lean_is_exclusive(x_46); -if (x_77 == 0) -{ -return x_46; +lean_dec(x_9); +return x_18; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_46, 0); -x_79 = lean_ctor_get(x_46, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_46); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; -} -} +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_6); +lean_dec(x_4); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +lean_dec(x_15); +x_21 = lean_box(0); +x_22 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_5, x_21); +x_23 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_22, x_21); +x_24 = l_Lean_MessageData_ofList(x_23); +x_25 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_7, x_9, x_10, x_11, x_12, x_19); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(x_1, x_2, x_20, x_9, x_10, x_11, x_12, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2; +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_29); +x_36 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_indentD(x_27); +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_indentD(x_32); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_25); +x_45 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_44, x_9, x_10, x_11, x_12, x_33); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +return x_45; } else { -lean_object* x_81; +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_45, 0); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_inc(x_47); lean_dec(x_45); -lean_dec(x_34); -lean_dec(x_31); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_9); -x_18 = x_81; -x_19 = x_43; -goto block_24; +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } else { -uint8_t x_82; -lean_dec(x_34); -lean_dec(x_31); -lean_free_object(x_26); -lean_dec(x_30); +uint8_t x_50; lean_dec(x_29); -lean_free_object(x_9); -lean_dec(x_28); -lean_dec(x_13); +lean_dec(x_27); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_82 = !lean_is_exclusive(x_41); -if (x_82 == 0) +lean_dec(x_9); +x_50 = !lean_is_exclusive(x_31); +if (x_50 == 0) { -return x_41; +return x_31; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_41, 0); -x_84 = lean_ctor_get(x_41, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_41); -x_85 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -return x_85; +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_31, 0); +x_52 = lean_ctor_get(x_31, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_31); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } else { -lean_object* x_86; uint8_t x_87; -x_86 = lean_ctor_get(x_37, 1); -lean_inc(x_86); -lean_dec(x_37); -x_87 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_28, x_34); -if (x_87 == 0) +uint8_t x_54; +lean_dec(x_27); +lean_dec(x_20); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_54 = !lean_is_exclusive(x_28); +if (x_54 == 0) { -lean_object* x_88; -lean_dec(x_31); -x_88 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_88, 0, x_9); -x_18 = x_88; -x_19 = x_86; -goto block_24; +return x_28; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; -lean_dec(x_29); -x_89 = lean_box(0); -lean_inc(x_31); -x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_28, x_31, x_89); -x_91 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_89); -x_92 = 1; -x_93 = lean_box(x_92); -lean_ctor_set(x_26, 1, x_91); -lean_ctor_set(x_26, 0, x_93); -lean_ctor_set(x_9, 0, x_90); -x_94 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_94, 0, x_9); -x_18 = x_94; -x_19 = x_86; -goto block_24; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_28, 0); +x_56 = lean_ctor_get(x_28, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_28); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} } } } else { -uint8_t x_95; -lean_dec(x_34); -lean_dec(x_31); -lean_free_object(x_26); -lean_dec(x_30); -lean_dec(x_29); -lean_free_object(x_9); -lean_dec(x_28); -lean_dec(x_13); +uint8_t x_58; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_95 = !lean_is_exclusive(x_37); -if (x_95 == 0) +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_14); +if (x_58 == 0) +{ +return x_14; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_14, 0); +x_60 = lean_ctor_get(x_14, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_14); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematch", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("pattern", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3() { +_start: { -return x_37; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1; +x_3 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4() { +_start: { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_37, 0); -x_97 = lean_ctor_get(x_37, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_37); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -return x_98; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} } +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_99; -lean_dec(x_29); -x_99 = !lean_is_exclusive(x_36); -if (x_99 == 0) +lean_object* x_11; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_11 = l_Lean_Meta_Grind_NormalizePattern_main(x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; -x_100 = lean_ctor_get(x_36, 0); -lean_dec(x_100); -x_101 = lean_box(0); -lean_inc(x_2); -x_102 = lean_array_mk(x_2); -x_103 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_104 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_101); -lean_ctor_set(x_104, 2, x_102); -x_105 = l_Lean_CollectFVars_main(x_34, x_104); -x_106 = lean_ctor_get(x_105, 2); -lean_inc(x_106); -lean_dec(x_105); -x_107 = lean_array_get_size(x_106); -x_108 = lean_unsigned_to_nat(0u); -x_109 = lean_nat_dec_lt(x_108, x_107); -x_110 = lean_box(0); -x_111 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_110); -if (x_109 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) { -uint8_t x_112; lean_object* x_113; -lean_dec(x_107); -lean_dec(x_106); -x_112 = 1; -x_113 = lean_box(x_112); -lean_ctor_set(x_26, 1, x_111); -lean_ctor_set(x_26, 0, x_113); -lean_ctor_set(x_36, 0, x_9); -x_18 = x_36; -x_19 = x_35; -goto block_24; +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_19 = lean_ctor_get(x_13, 0); +x_20 = lean_ctor_get(x_13, 1); +x_21 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_21, x_6, x_7, x_8, x_9, x_14); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_free_object(x_13); +lean_free_object(x_12); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_box(0); +x_27 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_20, x_2, x_16, x_19, x_1, x_26, x_6, x_7, x_8, x_9, x_25); +return x_27; } else { -uint8_t x_114; -x_114 = lean_nat_dec_le(x_107, x_107); -if (x_114 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_22); +if (x_28 == 0) { -uint8_t x_115; lean_object* x_116; -lean_dec(x_107); -lean_dec(x_106); -x_115 = 1; -x_116 = lean_box(x_115); -lean_ctor_set(x_26, 1, x_111); -lean_ctor_set(x_26, 0, x_116); -lean_ctor_set(x_36, 0, x_9); -x_18 = x_36; -x_19 = x_35; -goto block_24; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_29 = lean_ctor_get(x_22, 1); +x_30 = lean_ctor_get(x_22, 0); +lean_dec(x_30); +lean_inc(x_4); +x_31 = l_Lean_MessageData_ofConst(x_4); +x_32 = l_Lean_Meta_Grind_ppPattern___closed__4; +lean_ctor_set_tag(x_22, 7); +lean_ctor_set(x_22, 1, x_31); +lean_ctor_set(x_22, 0, x_32); +x_33 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +lean_ctor_set_tag(x_13, 7); +lean_ctor_set(x_13, 1, x_33); +lean_ctor_set(x_13, 0, x_22); +x_34 = lean_box(0); +lean_inc(x_16); +x_35 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_16, x_34); +x_36 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_35, x_34); +x_37 = l_Lean_MessageData_ofList(x_36); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_37); +lean_ctor_set(x_12, 0, x_13); +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_12); +lean_ctor_set(x_38, 1, x_32); +x_39 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_21, x_38, x_6, x_7, x_8, x_9, x_29); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_20, x_2, x_16, x_19, x_1, x_40, x_6, x_7, x_8, x_9, x_41); +lean_dec(x_40); +return x_42; } else { -size_t x_117; size_t x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; -x_117 = 0; -x_118 = lean_usize_of_nat(x_107); -lean_dec(x_107); -x_119 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_106, x_117, x_118, x_28); -lean_dec(x_106); -x_120 = 1; -x_121 = lean_box(x_120); -lean_ctor_set(x_26, 1, x_111); -lean_ctor_set(x_26, 0, x_121); -lean_ctor_set(x_9, 0, x_119); -lean_ctor_set(x_36, 0, x_9); -x_18 = x_36; -x_19 = x_35; -goto block_24; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_43 = lean_ctor_get(x_22, 1); +lean_inc(x_43); +lean_dec(x_22); +lean_inc(x_4); +x_44 = l_Lean_MessageData_ofConst(x_4); +x_45 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +x_47 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +lean_ctor_set_tag(x_13, 7); +lean_ctor_set(x_13, 1, x_47); +lean_ctor_set(x_13, 0, x_46); +x_48 = lean_box(0); +lean_inc(x_16); +x_49 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_16, x_48); +x_50 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_49, x_48); +x_51 = l_Lean_MessageData_ofList(x_50); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_51); +lean_ctor_set(x_12, 0, x_13); +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_12); +lean_ctor_set(x_52, 1, x_45); +x_53 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_21, x_52, x_6, x_7, x_8, x_9, x_43); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_20, x_2, x_16, x_19, x_1, x_54, x_6, x_7, x_8, x_9, x_55); +lean_dec(x_54); +return x_56; } } } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; -lean_dec(x_36); -x_122 = lean_box(0); -lean_inc(x_2); -x_123 = lean_array_mk(x_2); -x_124 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_125 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_122); -lean_ctor_set(x_125, 2, x_123); -x_126 = l_Lean_CollectFVars_main(x_34, x_125); -x_127 = lean_ctor_get(x_126, 2); -lean_inc(x_127); -lean_dec(x_126); -x_128 = lean_array_get_size(x_127); -x_129 = lean_unsigned_to_nat(0u); -x_130 = lean_nat_dec_lt(x_129, x_128); -x_131 = lean_box(0); -x_132 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_30, x_31, x_131); -if (x_130 == 0) +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_57 = lean_ctor_get(x_13, 0); +x_58 = lean_ctor_get(x_13, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_13); +x_59 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_60 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_59, x_6, x_7, x_8, x_9, x_14); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_unbox(x_61); +lean_dec(x_61); +if (x_62 == 0) { -uint8_t x_133; lean_object* x_134; lean_object* x_135; -lean_dec(x_128); -lean_dec(x_127); -x_133 = 1; -x_134 = lean_box(x_133); -lean_ctor_set(x_26, 1, x_132); -lean_ctor_set(x_26, 0, x_134); -x_135 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_135, 0, x_9); -x_18 = x_135; -x_19 = x_35; -goto block_24; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_free_object(x_12); +x_63 = lean_ctor_get(x_60, 1); +lean_inc(x_63); +lean_dec(x_60); +x_64 = lean_box(0); +x_65 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_58, x_2, x_16, x_57, x_1, x_64, x_6, x_7, x_8, x_9, x_63); +return x_65; } else { -uint8_t x_136; -x_136 = lean_nat_dec_le(x_128, x_128); -if (x_136 == 0) -{ -uint8_t x_137; lean_object* x_138; lean_object* x_139; -lean_dec(x_128); -lean_dec(x_127); -x_137 = 1; -x_138 = lean_box(x_137); -lean_ctor_set(x_26, 1, x_132); -lean_ctor_set(x_26, 0, x_138); -x_139 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_139, 0, x_9); -x_18 = x_139; -x_19 = x_35; -goto block_24; +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_67 = x_60; +} else { + lean_dec_ref(x_60); + x_67 = lean_box(0); +} +lean_inc(x_4); +x_68 = l_Lean_MessageData_ofConst(x_4); +x_69 = l_Lean_Meta_Grind_ppPattern___closed__4; +if (lean_is_scalar(x_67)) { + x_70 = lean_alloc_ctor(7, 2, 0); +} else { + x_70 = x_67; + lean_ctor_set_tag(x_70, 7); +} +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_68); +x_71 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_box(0); +lean_inc(x_16); +x_74 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_16, x_73); +x_75 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_74, x_73); +x_76 = l_Lean_MessageData_ofList(x_75); +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_76); +lean_ctor_set(x_12, 0, x_72); +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_12); +lean_ctor_set(x_77, 1, x_69); +x_78 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_59, x_77, x_6, x_7, x_8, x_9, x_66); +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_78, 1); +lean_inc(x_80); +lean_dec(x_78); +x_81 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_58, x_2, x_16, x_57, x_1, x_79, x_6, x_7, x_8, x_9, x_80); +lean_dec(x_79); +return x_81; +} +} } else { -size_t x_140; size_t x_141; lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; -x_140 = 0; -x_141 = lean_usize_of_nat(x_128); -lean_dec(x_128); -x_142 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_127, x_140, x_141, x_28); -lean_dec(x_127); -x_143 = 1; -x_144 = lean_box(x_143); -lean_ctor_set(x_26, 1, x_132); -lean_ctor_set(x_26, 0, x_144); -lean_ctor_set(x_9, 0, x_142); -x_145 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_145, 0, x_9); -x_18 = x_145; -x_19 = x_35; -goto block_24; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_82 = lean_ctor_get(x_12, 0); +lean_inc(x_82); +lean_dec(x_12); +x_83 = lean_ctor_get(x_13, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_13, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_85 = x_13; +} else { + lean_dec_ref(x_13); + x_85 = lean_box(0); +} +x_86 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_87 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_86, x_6, x_7, x_8, x_9, x_14); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_unbox(x_88); +lean_dec(x_88); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_85); +x_90 = lean_ctor_get(x_87, 1); +lean_inc(x_90); +lean_dec(x_87); +x_91 = lean_box(0); +x_92 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_84, x_2, x_82, x_83, x_1, x_91, x_6, x_7, x_8, x_9, x_90); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_93 = lean_ctor_get(x_87, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_94 = x_87; +} else { + lean_dec_ref(x_87); + x_94 = lean_box(0); +} +lean_inc(x_4); +x_95 = l_Lean_MessageData_ofConst(x_4); +x_96 = l_Lean_Meta_Grind_ppPattern___closed__4; +if (lean_is_scalar(x_94)) { + x_97 = lean_alloc_ctor(7, 2, 0); +} else { + x_97 = x_94; + lean_ctor_set_tag(x_97, 7); } +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +if (lean_is_scalar(x_85)) { + x_99 = lean_alloc_ctor(7, 2, 0); +} else { + x_99 = x_85; + lean_ctor_set_tag(x_99, 7); } +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = lean_box(0); +lean_inc(x_82); +x_101 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_82, x_100); +x_102 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_101, x_100); +x_103 = l_Lean_MessageData_ofList(x_102); +x_104 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_104, 0, x_99); +lean_ctor_set(x_104, 1, x_103); +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_96); +x_106 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_86, x_105, x_6, x_7, x_8, x_9, x_93); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_4, x_3, x_84, x_2, x_82, x_83, x_1, x_107, x_6, x_7, x_8, x_9, x_108); +lean_dec(x_107); +return x_109; } } } else { -uint8_t x_146; -lean_dec(x_31); -lean_free_object(x_26); -lean_dec(x_30); -lean_dec(x_29); -lean_free_object(x_9); -lean_dec(x_28); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_110; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_146 = !lean_is_exclusive(x_33); -if (x_146 == 0) +lean_dec(x_1); +x_110 = !lean_is_exclusive(x_11); +if (x_110 == 0) { -return x_33; +return x_11; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_33, 0); -x_148 = lean_ctor_get(x_33, 1); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_33); -x_149 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -return x_149; +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_11, 0); +x_112 = lean_ctor_get(x_11, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_11); +x_113 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +return x_113; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -uint8_t x_150; -lean_dec(x_31); -lean_dec(x_17); -x_150 = !lean_is_exclusive(x_32); -if (x_150 == 0) +lean_object* x_7; +x_7 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_151; -x_151 = lean_ctor_get(x_32, 0); -lean_dec(x_151); -lean_ctor_set(x_32, 0, x_9); -x_18 = x_32; -x_19 = x_14; -goto block_24; +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_13; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_152; -lean_dec(x_32); -x_152 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_152, 0, x_9); -x_18 = x_152; -x_19 = x_14; -goto block_24; +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +return x_14; } } +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_ctor_get(x_2, 5); +x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_5); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_8); +lean_ctor_set_tag(x_6, 1); +lean_ctor_set(x_6, 0, x_9); +return x_6; } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_153 = lean_ctor_get(x_9, 0); -x_154 = lean_ctor_get(x_26, 0); -x_155 = lean_ctor_get(x_26, 1); -lean_inc(x_155); -lean_inc(x_154); -lean_dec(x_26); -x_156 = l_Lean_Expr_fvarId_x21(x_17); -x_157 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_155, x_156); -if (lean_obj_tag(x_157) == 0) -{ -lean_object* x_158; -lean_inc(x_13); -lean_inc(x_12); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); lean_inc(x_11); lean_inc(x_10); -x_158 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_158) == 0) +lean_dec(x_6); +lean_inc(x_5); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1() { +_start: { -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_158, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_158, 1); -lean_inc(x_160); -lean_dec(x_158); -x_161 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_153, x_156); -if (lean_obj_tag(x_161) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is not a theorem", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2() { +_start: { -lean_object* x_162; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_159); -x_162 = l_Lean_Meta_isProp(x_159, x_10, x_11, x_12, x_13, x_160); -if (lean_obj_tag(x_162) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_163; uint8_t x_164; -x_163 = lean_ctor_get(x_162, 0); -lean_inc(x_163); -x_164 = lean_unbox(x_163); -lean_dec(x_163); -if (x_164 == 0) +lean_object* x_5; +lean_inc(x_1); +x_5 = l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_165; lean_object* x_166; -x_165 = lean_ctor_get(x_162, 1); -lean_inc(x_165); -lean_dec(x_162); -lean_inc(x_10); -lean_inc(x_156); -x_166 = l_Lean_FVarId_getDecl(x_156, x_10, x_11, x_12, x_13, x_165); -if (lean_obj_tag(x_166) == 0) +lean_object* x_6; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +if (lean_obj_tag(x_6) == 2) { -lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; -x_167 = lean_ctor_get(x_166, 0); -lean_inc(x_167); -x_168 = lean_ctor_get(x_166, 1); -lean_inc(x_168); -lean_dec(x_166); -x_169 = l_Lean_LocalDecl_binderInfo(x_167); -lean_dec(x_167); -x_170 = lean_box(x_169); -if (lean_obj_tag(x_170) == 3) +uint8_t x_7; +x_7 = !lean_is_exclusive(x_5); +if (x_7 == 0) { -lean_object* x_171; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_8 = lean_ctor_get(x_5, 0); +lean_dec(x_8); +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_inc(x_159); -lean_inc(x_153); -lean_inc(x_4); -x_171 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_153, x_159, x_10, x_11, x_12, x_13, x_168); -if (lean_obj_tag(x_171) == 0) +lean_dec(x_9); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_box(0); +x_13 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_11, x_12); +x_14 = l_Lean_Expr_const___override(x_1, x_13); +lean_ctor_set(x_5, 0, x_14); +return x_5; +} +else { -lean_object* x_172; uint8_t x_173; -x_172 = lean_ctor_get(x_171, 0); -lean_inc(x_172); -x_173 = lean_unbox(x_172); -lean_dec(x_172); -if (x_173 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_15 = lean_ctor_get(x_5, 1); +lean_inc(x_15); +lean_dec(x_5); +x_16 = lean_ctor_get(x_6, 0); +lean_inc(x_16); +lean_dec(x_6); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_box(0); +x_20 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_18, x_19); +x_21 = l_Lean_Expr_const___override(x_1, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_15); +return x_22; +} +} +else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; -lean_dec(x_159); -lean_dec(x_156); -x_174 = lean_ctor_get(x_171, 1); -lean_inc(x_174); -lean_dec(x_171); -x_175 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_175, 0, x_154); -lean_ctor_set(x_175, 1, x_155); -lean_ctor_set(x_9, 1, x_175); -x_176 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_176, 0, x_9); -x_18 = x_176; -x_19 = x_174; -goto block_24; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_6); +x_23 = lean_ctor_get(x_5, 1); +lean_inc(x_23); +lean_dec(x_5); +x_24 = l_Lean_MessageData_ofName(x_1); +x_25 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1(x_28, x_2, x_3, x_23); +return x_29; +} } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; lean_object* x_189; -lean_dec(x_154); -x_177 = lean_ctor_get(x_171, 1); -lean_inc(x_177); -lean_dec(x_171); -x_178 = lean_box(0); -lean_inc(x_156); -x_179 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_178); -x_180 = lean_box(0); -lean_inc(x_2); -x_181 = lean_array_mk(x_2); -x_182 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_183 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_183, 0, x_182); -lean_ctor_set(x_183, 1, x_180); -lean_ctor_set(x_183, 2, x_181); -x_184 = l_Lean_CollectFVars_main(x_159, x_183); -x_185 = lean_ctor_get(x_184, 2); -lean_inc(x_185); -lean_dec(x_184); -x_186 = lean_array_get_size(x_185); -x_187 = lean_unsigned_to_nat(0u); -x_188 = lean_nat_dec_lt(x_187, x_186); -x_189 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_178); -if (x_188 == 0) +uint8_t x_30; +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_5); +if (x_30 == 0) { -uint8_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -lean_dec(x_186); -lean_dec(x_185); -x_190 = 1; -x_191 = lean_box(x_190); -x_192 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_189); -lean_ctor_set(x_9, 1, x_192); -lean_ctor_set(x_9, 0, x_179); -x_193 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_193, 0, x_9); -x_18 = x_193; -x_19 = x_177; -goto block_24; +return x_5; } else { -uint8_t x_194; -x_194 = lean_nat_dec_le(x_186, x_186); -if (x_194 == 0) -{ -uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -lean_dec(x_186); -lean_dec(x_185); -x_195 = 1; -x_196 = lean_box(x_195); -x_197 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_197, 0, x_196); -lean_ctor_set(x_197, 1, x_189); -lean_ctor_set(x_9, 1, x_197); -lean_ctor_set(x_9, 0, x_179); -x_198 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_198, 0, x_9); -x_18 = x_198; -x_19 = x_177; -goto block_24; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_5, 0); +x_32 = lean_ctor_get(x_5, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_5); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } -else +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -size_t x_199; size_t x_200; lean_object* x_201; uint8_t x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_199 = 0; -x_200 = lean_usize_of_nat(x_186); -lean_dec(x_186); -x_201 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_185, x_199, x_200, x_179); -lean_dec(x_185); -x_202 = 1; -x_203 = lean_box(x_202); -x_204 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_204, 0, x_203); -lean_ctor_set(x_204, 1, x_189); -lean_ctor_set(x_9, 1, x_204); -lean_ctor_set(x_9, 0, x_201); -x_205 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_205, 0, x_9); -x_18 = x_205; -x_19 = x_177; -goto block_24; +lean_object* x_5; +x_5 = l_Lean_throwError___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_1); +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(x_1, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_1); +x_13 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_14 = l_Lean_Meta_Grind_mkEMatchTheoremCore(x_12, x_13, x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; } else { -lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_159); -lean_dec(x_156); -lean_dec(x_155); -lean_dec(x_154); -lean_free_object(x_9); -lean_dec(x_153); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_206 = lean_ctor_get(x_171, 0); -lean_inc(x_206); -x_207 = lean_ctor_get(x_171, 1); -lean_inc(x_207); -if (lean_is_exclusive(x_171)) { - lean_ctor_release(x_171, 0); - lean_ctor_release(x_171, 1); - x_208 = x_171; -} else { - lean_dec_ref(x_171); - x_208 = lean_box(0); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +return x_9; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } -if (lean_is_scalar(x_208)) { - x_209 = lean_alloc_ctor(1, 2, 0); -} else { - x_209 = x_208; } -lean_ctor_set(x_209, 0, x_206); -lean_ctor_set(x_209, 1, x_207); -return x_209; } } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_4, 5); +x_8 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set_tag(x_8, 1); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} else { -lean_object* x_210; lean_object* x_211; -lean_dec(x_170); -lean_dec(x_159); -lean_dec(x_156); -x_210 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_210, 0, x_154); -lean_ctor_set(x_210, 1, x_155); -lean_ctor_set(x_9, 1, x_210); -x_211 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_211, 0, x_9); -x_18 = x_211; -x_19 = x_168; -goto block_24; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +lean_inc(x_7); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1(uint8_t x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; -lean_dec(x_159); -lean_dec(x_156); -lean_dec(x_155); -lean_dec(x_154); -lean_free_object(x_9); -lean_dec(x_153); -lean_dec(x_13); +if (x_1 == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_4, 1); +x_12 = lean_ctor_get(x_4, 0); lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_212 = lean_ctor_get(x_166, 0); -lean_inc(x_212); -x_213 = lean_ctor_get(x_166, 1); -lean_inc(x_213); -if (lean_is_exclusive(x_166)) { - lean_ctor_release(x_166, 0); - lean_ctor_release(x_166, 1); - x_214 = x_166; -} else { - lean_dec_ref(x_166); - x_214 = lean_box(0); -} -if (lean_is_scalar(x_214)) { - x_215 = lean_alloc_ctor(1, 2, 0); -} else { - x_215 = x_214; +x_13 = l_Lean_Meta_Grind_preprocessPattern(x_11, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_array_get_size(x_3); +x_17 = lean_expr_abstract(x_15, x_3); +lean_dec(x_15); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_4, 1, x_19); +lean_ctor_set(x_4, 0, x_16); +lean_ctor_set(x_13, 0, x_4); +return x_13; } -lean_ctor_set(x_215, 0, x_212); -lean_ctor_set(x_215, 1, x_213); -return x_215; +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_array_get_size(x_3); +x_23 = lean_expr_abstract(x_20, x_3); +lean_dec(x_20); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_4, 1, x_25); +lean_ctor_set(x_4, 0, x_22); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_4); +lean_ctor_set(x_26, 1, x_21); +return x_26; } } else { -lean_object* x_216; uint8_t x_217; -x_216 = lean_ctor_get(x_162, 1); -lean_inc(x_216); -lean_dec(x_162); -x_217 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_153, x_159); -if (x_217 == 0) +uint8_t x_27; +lean_free_object(x_4); +x_27 = !lean_is_exclusive(x_13); +if (x_27 == 0) { -lean_object* x_218; lean_object* x_219; -lean_dec(x_156); -x_218 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_218, 0, x_154); -lean_ctor_set(x_218, 1, x_155); -lean_ctor_set(x_9, 1, x_218); -x_219 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_219, 0, x_9); -x_18 = x_219; -x_19 = x_216; -goto block_24; +return x_13; } else { -lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; -lean_dec(x_154); -x_220 = lean_box(0); -lean_inc(x_156); -x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_153, x_156, x_220); -x_222 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_220); -x_223 = 1; -x_224 = lean_box(x_223); -x_225 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_225, 0, x_224); -lean_ctor_set(x_225, 1, x_222); -lean_ctor_set(x_9, 1, x_225); -lean_ctor_set(x_9, 0, x_221); -x_226 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_226, 0, x_9); -x_18 = x_226; -x_19 = x_216; -goto block_24; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_13, 0); +x_29 = lean_ctor_get(x_13, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_13); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } } else { -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; -lean_dec(x_159); -lean_dec(x_156); -lean_dec(x_155); -lean_dec(x_154); -lean_free_object(x_9); -lean_dec(x_153); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_4, 1); +lean_inc(x_31); lean_dec(x_4); -lean_dec(x_2); -x_227 = lean_ctor_get(x_162, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_162, 1); -lean_inc(x_228); -if (lean_is_exclusive(x_162)) { - lean_ctor_release(x_162, 0); - lean_ctor_release(x_162, 1); - x_229 = x_162; +x_32 = l_Lean_Meta_Grind_preprocessPattern(x_31, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_35 = x_32; } else { - lean_dec_ref(x_162); - x_229 = lean_box(0); + lean_dec_ref(x_32); + x_35 = lean_box(0); } -if (lean_is_scalar(x_229)) { - x_230 = lean_alloc_ctor(1, 2, 0); +x_36 = lean_array_get_size(x_3); +x_37 = lean_expr_abstract(x_33, x_3); +lean_dec(x_33); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_36); +lean_ctor_set(x_40, 1, x_39); +if (lean_is_scalar(x_35)) { + x_41 = lean_alloc_ctor(0, 2, 0); } else { - x_230 = x_229; -} -lean_ctor_set(x_230, 0, x_227); -lean_ctor_set(x_230, 1, x_228); -return x_230; + x_41 = x_35; } +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_34); +return x_41; } else { -lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; uint8_t x_240; lean_object* x_241; lean_object* x_242; -lean_dec(x_154); -if (lean_is_exclusive(x_161)) { - lean_ctor_release(x_161, 0); - x_231 = x_161; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_32, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_32, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_44 = x_32; } else { - lean_dec_ref(x_161); - x_231 = lean_box(0); + lean_dec_ref(x_32); + x_44 = lean_box(0); } -x_232 = lean_box(0); -lean_inc(x_2); -x_233 = lean_array_mk(x_2); -x_234 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_235 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_232); -lean_ctor_set(x_235, 2, x_233); -x_236 = l_Lean_CollectFVars_main(x_159, x_235); -x_237 = lean_ctor_get(x_236, 2); -lean_inc(x_237); -lean_dec(x_236); -x_238 = lean_array_get_size(x_237); -x_239 = lean_unsigned_to_nat(0u); -x_240 = lean_nat_dec_lt(x_239, x_238); -x_241 = lean_box(0); -x_242 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_155, x_156, x_241); -if (x_240 == 0) -{ -uint8_t x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -lean_dec(x_238); -lean_dec(x_237); -x_243 = 1; -x_244 = lean_box(x_243); -x_245 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_245, 0, x_244); -lean_ctor_set(x_245, 1, x_242); -lean_ctor_set(x_9, 1, x_245); -if (lean_is_scalar(x_231)) { - x_246 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(1, 2, 0); } else { - x_246 = x_231; + x_45 = x_44; +} +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +return x_45; +} } -lean_ctor_set(x_246, 0, x_9); -x_18 = x_246; -x_19 = x_160; -goto block_24; } else { -uint8_t x_247; -x_247 = lean_nat_dec_le(x_238, x_238); -if (x_247 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_4); +if (x_46 == 0) { -uint8_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -lean_dec(x_238); -lean_dec(x_237); -x_248 = 1; -x_249 = lean_box(x_248); -x_250 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_250, 0, x_249); -lean_ctor_set(x_250, 1, x_242); -lean_ctor_set(x_9, 1, x_250); -if (lean_is_scalar(x_231)) { - x_251 = lean_alloc_ctor(1, 1, 0); -} else { - x_251 = x_231; -} -lean_ctor_set(x_251, 0, x_9); -x_18 = x_251; -x_19 = x_160; -goto block_24; +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_4, 0); +x_48 = lean_ctor_get(x_4, 1); +lean_dec(x_48); +x_49 = l_Lean_Meta_Grind_preprocessPattern(x_47, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_51 = lean_ctor_get(x_49, 0); +x_52 = lean_array_get_size(x_3); +x_53 = lean_expr_abstract(x_51, x_3); +lean_dec(x_51); +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_4, 1, x_55); +lean_ctor_set(x_4, 0, x_52); +lean_ctor_set(x_49, 0, x_4); +return x_49; } else { -size_t x_252; size_t x_253; lean_object* x_254; uint8_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_252 = 0; -x_253 = lean_usize_of_nat(x_238); -lean_dec(x_238); -x_254 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_237, x_252, x_253, x_153); -lean_dec(x_237); -x_255 = 1; -x_256 = lean_box(x_255); -x_257 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_257, 0, x_256); -lean_ctor_set(x_257, 1, x_242); -lean_ctor_set(x_9, 1, x_257); -lean_ctor_set(x_9, 0, x_254); -if (lean_is_scalar(x_231)) { - x_258 = lean_alloc_ctor(1, 1, 0); -} else { - x_258 = x_231; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_56 = lean_ctor_get(x_49, 0); +x_57 = lean_ctor_get(x_49, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_49); +x_58 = lean_array_get_size(x_3); +x_59 = lean_expr_abstract(x_56, x_3); +lean_dec(x_56); +x_60 = lean_box(0); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_4, 1, x_61); +lean_ctor_set(x_4, 0, x_58); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_4); +lean_ctor_set(x_62, 1, x_57); +return x_62; } -lean_ctor_set(x_258, 0, x_9); -x_18 = x_258; -x_19 = x_160; -goto block_24; } +else +{ +uint8_t x_63; +lean_free_object(x_4); +x_63 = !lean_is_exclusive(x_49); +if (x_63 == 0) +{ +return x_49; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_49, 0); +x_65 = lean_ctor_get(x_49, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_49); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } else { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -lean_dec(x_156); -lean_dec(x_155); -lean_dec(x_154); -lean_free_object(x_9); -lean_dec(x_153); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_4, 0); +lean_inc(x_67); lean_dec(x_4); -lean_dec(x_2); -x_259 = lean_ctor_get(x_158, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_158, 1); -lean_inc(x_260); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_261 = x_158; -} else { - lean_dec_ref(x_158); - x_261 = lean_box(0); -} -if (lean_is_scalar(x_261)) { - x_262 = lean_alloc_ctor(1, 2, 0); +x_68 = l_Lean_Meta_Grind_preprocessPattern(x_67, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_71 = x_68; } else { - x_262 = x_261; -} -lean_ctor_set(x_262, 0, x_259); -lean_ctor_set(x_262, 1, x_260); -return x_262; + lean_dec_ref(x_68); + x_71 = lean_box(0); +} +x_72 = lean_array_get_size(x_3); +x_73 = lean_expr_abstract(x_69, x_3); +lean_dec(x_69); +x_74 = lean_box(0); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_72); +lean_ctor_set(x_76, 1, x_75); +if (lean_is_scalar(x_71)) { + x_77 = lean_alloc_ctor(0, 2, 0); +} else { + x_77 = x_71; } +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_70); +return x_77; } else { -lean_object* x_263; lean_object* x_264; lean_object* x_265; -lean_dec(x_156); -lean_dec(x_17); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - x_263 = x_157; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_68, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_68, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_80 = x_68; } else { - lean_dec_ref(x_157); - x_263 = lean_box(0); + lean_dec_ref(x_68); + x_80 = lean_box(0); } -x_264 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_264, 0, x_154); -lean_ctor_set(x_264, 1, x_155); -lean_ctor_set(x_9, 1, x_264); -if (lean_is_scalar(x_263)) { - x_265 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); } else { - x_265 = x_263; + x_81 = x_80; } -lean_ctor_set(x_265, 0, x_9); -x_18 = x_265; -x_19 = x_14; -goto block_24; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } } -else -{ -lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_266 = lean_ctor_get(x_9, 1); -x_267 = lean_ctor_get(x_9, 0); -lean_inc(x_266); -lean_inc(x_267); -lean_dec(x_9); -x_268 = lean_ctor_get(x_266, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_266, 1); -lean_inc(x_269); -if (lean_is_exclusive(x_266)) { - lean_ctor_release(x_266, 0); - lean_ctor_release(x_266, 1); - x_270 = x_266; -} else { - lean_dec_ref(x_266); - x_270 = lean_box(0); } -x_271 = l_Lean_Expr_fvarId_x21(x_17); -x_272 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_269, x_271); -if (lean_obj_tag(x_272) == 0) +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_273; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_273 = lean_infer_type(x_17, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_273) == 0) +lean_object* x_9; lean_object* x_10; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_2); +lean_ctor_set(x_9, 1, x_3); +x_10 = lean_apply_6(x_1, x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1() { +_start: { -lean_object* x_274; lean_object* x_275; lean_object* x_276; -x_274 = lean_ctor_get(x_273, 0); -lean_inc(x_274); -x_275 = lean_ctor_get(x_273, 1); -lean_inc(x_275); -lean_dec(x_273); -x_276 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_267, x_271); -if (lean_obj_tag(x_276) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid E-matching equality theorem, conclusion must be an equality", 67, 67); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2() { +_start: { -lean_object* x_277; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_274); -x_277 = l_Lean_Meta_isProp(x_274, x_10, x_11, x_12, x_13, x_275); -if (lean_obj_tag(x_277) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_278; uint8_t x_279; -x_278 = lean_ctor_get(x_277, 0); -lean_inc(x_278); -x_279 = lean_unbox(x_278); -lean_dec(x_278); -if (x_279 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_9 = l_Lean_indentExpr(x_1); +x_10 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2; +x_11 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_13 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1(x_13, x_4, x_5, x_6, x_7, x_8); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_280; lean_object* x_281; -x_280 = lean_ctor_get(x_277, 1); -lean_inc(x_280); -lean_dec(x_277); -lean_inc(x_10); -lean_inc(x_271); -x_281 = l_Lean_FVarId_getDecl(x_271, x_10, x_11, x_12, x_13, x_280); -if (lean_obj_tag(x_281) == 0) +return x_14; +} +else { -lean_object* x_282; lean_object* x_283; uint8_t x_284; lean_object* x_285; -x_282 = lean_ctor_get(x_281, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_281, 1); -lean_inc(x_283); -lean_dec(x_281); -x_284 = l_Lean_LocalDecl_binderInfo(x_282); -lean_dec(x_282); -x_285 = lean_box(x_284); -if (lean_obj_tag(x_285) == 3) +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4(uint8_t x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_286; -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_274); -lean_inc(x_267); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_10 = lean_box(x_1); +x_11 = lean_box(x_2); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1___boxed), 9, 3); +lean_closure_set(x_12, 0, x_10); +lean_closure_set(x_12, 1, x_11); +lean_closure_set(x_12, 2, x_3); lean_inc(x_4); -x_286 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_4, x_267, x_274, x_10, x_11, x_12, x_13, x_283); -if (lean_obj_tag(x_286) == 0) -{ -lean_object* x_287; uint8_t x_288; -x_287 = lean_ctor_get(x_286, 0); -lean_inc(x_287); -x_288 = lean_unbox(x_287); -lean_dec(x_287); -if (x_288 == 0) +x_13 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_4, x_5, x_6, x_7, x_8, x_9); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Expr_cleanupAnnotations(x_14); +x_17 = l_Lean_Expr_isApp(x_16); +if (x_17 == 0) { -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; -lean_dec(x_274); -lean_dec(x_271); -x_289 = lean_ctor_get(x_286, 1); -lean_inc(x_289); -lean_dec(x_286); -if (lean_is_scalar(x_270)) { - x_290 = lean_alloc_ctor(0, 2, 0); -} else { - x_290 = x_270; -} -lean_ctor_set(x_290, 0, x_268); -lean_ctor_set(x_290, 1, x_269); -x_291 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_291, 0, x_267); -lean_ctor_set(x_291, 1, x_290); -x_292 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_292, 0, x_291); -x_18 = x_292; -x_19 = x_289; -goto block_24; +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +x_18 = lean_box(0); +x_19 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_18, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_19; } else { -lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; lean_object* x_305; -lean_dec(x_268); -x_293 = lean_ctor_get(x_286, 1); -lean_inc(x_293); -lean_dec(x_286); -x_294 = lean_box(0); -lean_inc(x_271); -x_295 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_294); -x_296 = lean_box(0); -lean_inc(x_2); -x_297 = lean_array_mk(x_2); -x_298 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_299 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_299, 1, x_296); -lean_ctor_set(x_299, 2, x_297); -x_300 = l_Lean_CollectFVars_main(x_274, x_299); -x_301 = lean_ctor_get(x_300, 2); -lean_inc(x_301); -lean_dec(x_300); -x_302 = lean_array_get_size(x_301); -x_303 = lean_unsigned_to_nat(0u); -x_304 = lean_nat_dec_lt(x_303, x_302); -x_305 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_294); -if (x_304 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Expr_appArg(x_16, lean_box(0)); +x_21 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); +x_22 = l_Lean_Expr_isApp(x_21); +if (x_22 == 0) { -uint8_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -lean_dec(x_302); -lean_dec(x_301); -x_306 = 1; -x_307 = lean_box(x_306); -if (lean_is_scalar(x_270)) { - x_308 = lean_alloc_ctor(0, 2, 0); -} else { - x_308 = x_270; +lean_object* x_23; lean_object* x_24; +lean_dec(x_21); +lean_dec(x_20); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_23, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_24; } -lean_ctor_set(x_308, 0, x_307); -lean_ctor_set(x_308, 1, x_305); -x_309 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_309, 0, x_295); -lean_ctor_set(x_309, 1, x_308); -x_310 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_310, 0, x_309); -x_18 = x_310; -x_19 = x_293; -goto block_24; +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = l_Lean_Expr_appArg(x_21, lean_box(0)); +x_26 = l_Lean_Expr_appFnCleanup(x_21, lean_box(0)); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__6; +x_28 = l_Lean_Expr_isConstOf(x_26, x_27); +if (x_28 == 0) +{ +uint8_t x_29; +x_29 = l_Lean_Expr_isApp(x_26); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_20); +x_30 = lean_box(0); +x_31 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_30, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_31; } else { -uint8_t x_311; -x_311 = lean_nat_dec_le(x_302, x_302); -if (x_311 == 0) +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = l_Lean_Expr_appArg(x_26, lean_box(0)); +x_33 = l_Lean_Expr_appFnCleanup(x_26, lean_box(0)); +x_34 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__2; +x_35 = l_Lean_Expr_isConstOf(x_33, x_34); +if (x_35 == 0) { -uint8_t x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; -lean_dec(x_302); -lean_dec(x_301); -x_312 = 1; -x_313 = lean_box(x_312); -if (lean_is_scalar(x_270)) { - x_314 = lean_alloc_ctor(0, 2, 0); -} else { - x_314 = x_270; +uint8_t x_36; +lean_dec(x_25); +x_36 = l_Lean_Expr_isApp(x_33); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_20); +x_37 = lean_box(0); +x_38 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_37, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_38; } -lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_305); -x_315 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_315, 0, x_295); -lean_ctor_set(x_315, 1, x_314); -x_316 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_316, 0, x_315); -x_18 = x_316; -x_19 = x_293; -goto block_24; +else +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = l_Lean_Expr_appFnCleanup(x_33, lean_box(0)); +x_40 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames___closed__4; +x_41 = l_Lean_Expr_isConstOf(x_39, x_40); +lean_dec(x_39); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_32); +lean_dec(x_20); +x_42 = lean_box(0); +x_43 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_4, x_12, x_42, x_5, x_6, x_7, x_8, x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_12); +return x_43; } else { -size_t x_317; size_t x_318; lean_object* x_319; uint8_t x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; -x_317 = 0; -x_318 = lean_usize_of_nat(x_302); -lean_dec(x_302); -x_319 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_4, x_301, x_317, x_318, x_295); -lean_dec(x_301); -x_320 = 1; -x_321 = lean_box(x_320); -if (lean_is_scalar(x_270)) { - x_322 = lean_alloc_ctor(0, 2, 0); -} else { - x_322 = x_270; +lean_object* x_44; +lean_dec(x_4); +x_44 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(x_12, x_32, x_20, x_5, x_6, x_7, x_8, x_15); +return x_44; } -lean_ctor_set(x_322, 0, x_321); -lean_ctor_set(x_322, 1, x_305); -x_323 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_323, 0, x_319); -lean_ctor_set(x_323, 1, x_322); -x_324 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_324, 0, x_323); -x_18 = x_324; -x_19 = x_293; -goto block_24; } } +else +{ +lean_object* x_45; +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_4); +x_45 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(x_12, x_25, x_20, x_5, x_6, x_7, x_8, x_15); +return x_45; +} } } else { -lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; -lean_dec(x_274); -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_46; +lean_dec(x_26); lean_dec(x_4); -lean_dec(x_2); -x_325 = lean_ctor_get(x_286, 0); -lean_inc(x_325); -x_326 = lean_ctor_get(x_286, 1); -lean_inc(x_326); -if (lean_is_exclusive(x_286)) { - lean_ctor_release(x_286, 0); - lean_ctor_release(x_286, 1); - x_327 = x_286; -} else { - lean_dec_ref(x_286); - x_327 = lean_box(0); +x_46 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__2(x_12, x_25, x_20, x_5, x_6, x_7, x_8, x_15); +return x_46; } -if (lean_is_scalar(x_327)) { - x_328 = lean_alloc_ctor(1, 2, 0); -} else { - x_328 = x_327; } -lean_ctor_set(x_328, 0, x_325); -lean_ctor_set(x_328, 1, x_326); -return x_328; } } +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_3); +x_11 = lean_infer_type(x_3, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_box(x_5); +x_15 = lean_box(x_4); +x_16 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4___boxed), 9, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = 0; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_18 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_12, x_16, x_17, x_6, x_7, x_8, x_9, x_13); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = l_Lean_Meta_Grind_mkEMatchTheoremCore(x_1, x_2, x_21, x_3, x_22, x_6, x_7, x_8, x_9, x_20); +return x_23; +} +else +{ +uint8_t x_24; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +return x_18; +} else { -lean_object* x_329; lean_object* x_330; lean_object* x_331; -lean_dec(x_285); -lean_dec(x_274); -lean_dec(x_271); -if (lean_is_scalar(x_270)) { - x_329 = lean_alloc_ctor(0, 2, 0); -} else { - x_329 = x_270; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_18, 0); +x_26 = lean_ctor_get(x_18, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_18); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } -lean_ctor_set(x_329, 0, x_268); -lean_ctor_set(x_329, 1, x_269); -x_330 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_330, 0, x_267); -lean_ctor_set(x_330, 1, x_329); -x_331 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_331, 0, x_330); -x_18 = x_331; -x_19 = x_283; -goto block_24; } } else { -lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; -lean_dec(x_274); -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_28; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_11); +if (x_28 == 0) +{ +return x_11; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_11, 0); +x_30 = lean_ctor_get(x_11, 1); +lean_inc(x_30); +lean_inc(x_29); lean_dec(x_11); -lean_dec(x_10); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_throwError___at_Lean_Meta_Grind_mkEMatchEqTheoremCore___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; uint8_t x_11; lean_object* x_12; +x_10 = lean_unbox(x_1); +lean_dec(x_1); +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__1(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; uint8_t x_11; lean_object* x_12; +x_10 = lean_unbox(x_1); +lean_dec(x_1); +x_11 = lean_unbox(x_2); lean_dec(x_2); -x_332 = lean_ctor_get(x_281, 0); -lean_inc(x_332); -x_333 = lean_ctor_get(x_281, 1); -lean_inc(x_333); -if (lean_is_exclusive(x_281)) { - lean_ctor_release(x_281, 0); - lean_ctor_release(x_281, 1); - x_334 = x_281; -} else { - lean_dec_ref(x_281); - x_334 = lean_box(0); +x_12 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__4(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } -if (lean_is_scalar(x_334)) { - x_335 = lean_alloc_ctor(1, 2, 0); -} else { - x_335 = x_334; } -lean_ctor_set(x_335, 0, x_332); -lean_ctor_set(x_335, 1, x_333); -return x_335; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheoremCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore(x_1, x_2, x_3, x_11, x_12, x_6, x_7, x_8, x_9, x_10); +return x_13; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_336; uint8_t x_337; -x_336 = lean_ctor_get(x_277, 1); -lean_inc(x_336); -lean_dec(x_277); -x_337 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_4, x_267, x_274); -if (x_337 == 0) +lean_object* x_9; +lean_inc(x_1); +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(x_1, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_338; lean_object* x_339; lean_object* x_340; -lean_dec(x_271); -if (lean_is_scalar(x_270)) { - x_338 = lean_alloc_ctor(0, 2, 0); -} else { - x_338 = x_270; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_1); +x_13 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_14 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore(x_12, x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; } -lean_ctor_set(x_338, 0, x_268); -lean_ctor_set(x_338, 1, x_269); -x_339 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_339, 0, x_267); -lean_ctor_set(x_339, 1, x_338); -x_340 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_340, 0, x_339); -x_18 = x_340; -x_19 = x_336; -goto block_24; +else +{ +uint8_t x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +return x_9; } else { -lean_object* x_341; lean_object* x_342; lean_object* x_343; uint8_t x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; -lean_dec(x_268); -x_341 = lean_box(0); -lean_inc(x_271); -x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_267, x_271, x_341); -x_343 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_341); -x_344 = 1; -x_345 = lean_box(x_344); -if (lean_is_scalar(x_270)) { - x_346 = lean_alloc_ctor(0, 2, 0); -} else { - x_346 = x_270; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } -lean_ctor_set(x_346, 0, x_345); -lean_ctor_set(x_346, 1, x_343); -x_347 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_347, 0, x_342); -lean_ctor_set(x_347, 1, x_346); -x_348 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_348, 0, x_347); -x_18 = x_348; -x_19 = x_336; -goto block_24; } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; -lean_dec(x_274); -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); +uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_unbox(x_2); lean_dec(x_2); -x_349 = lean_ctor_get(x_277, 0); -lean_inc(x_349); -x_350 = lean_ctor_get(x_277, 1); -lean_inc(x_350); -if (lean_is_exclusive(x_277)) { - lean_ctor_release(x_277, 0); - lean_ctor_release(x_277, 1); - x_351 = x_277; -} else { - lean_dec_ref(x_277); - x_351 = lean_box(0); -} -if (lean_is_scalar(x_351)) { - x_352 = lean_alloc_ctor(1, 2, 0); -} else { - x_352 = x_351; -} -lean_ctor_set(x_352, 0, x_349); -lean_ctor_set(x_352, 1, x_350); -return x_352; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; } } -else +static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1() { +_start: { -lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; uint8_t x_362; lean_object* x_363; lean_object* x_364; -lean_dec(x_268); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - x_353 = x_276; -} else { - lean_dec_ref(x_276); - x_353 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; } -x_354 = lean_box(0); -lean_inc(x_2); -x_355 = lean_array_mk(x_2); -x_356 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_357 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_357, 0, x_356); -lean_ctor_set(x_357, 1, x_354); -lean_ctor_set(x_357, 2, x_355); -x_358 = l_Lean_CollectFVars_main(x_274, x_357); -x_359 = lean_ctor_get(x_358, 2); -lean_inc(x_359); -lean_dec(x_358); -x_360 = lean_array_get_size(x_359); -x_361 = lean_unsigned_to_nat(0u); -x_362 = lean_nat_dec_lt(x_361, x_360); -x_363 = lean_box(0); -x_364 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_269, x_271, x_363); -if (x_362 == 0) -{ -uint8_t x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; -lean_dec(x_360); -lean_dec(x_359); -x_365 = 1; -x_366 = lean_box(x_365); -if (lean_is_scalar(x_270)) { - x_367 = lean_alloc_ctor(0, 2, 0); -} else { - x_367 = x_270; } -lean_ctor_set(x_367, 0, x_366); -lean_ctor_set(x_367, 1, x_364); -x_368 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_368, 0, x_267); -lean_ctor_set(x_368, 1, x_367); -if (lean_is_scalar(x_353)) { - x_369 = lean_alloc_ctor(1, 1, 0); -} else { - x_369 = x_353; +static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 3, x_1); +lean_ctor_set(x_2, 4, x_1); +lean_ctor_set(x_2, 5, x_1); +return x_2; } -lean_ctor_set(x_369, 0, x_368); -x_18 = x_369; -x_19 = x_275; -goto block_24; } -else +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -uint8_t x_370; -x_370 = lean_nat_dec_le(x_360, x_360); -if (x_370 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_6, 6); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_st_ref_take(x_7, x_8); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) { -uint8_t x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -lean_dec(x_360); -lean_dec(x_359); -x_371 = 1; -x_372 = lean_box(x_371); -if (lean_is_scalar(x_270)) { - x_373 = lean_alloc_ctor(0, 2, 0); -} else { - x_373 = x_270; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 4); +lean_dec(x_15); +x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); +x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; +lean_ctor_set(x_11, 4, x_17); +lean_ctor_set(x_11, 0, x_16); +x_18 = lean_st_ref_set(x_7, x_11, x_12); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_st_ref_take(x_5, x_19); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_21, 1); +lean_dec(x_24); +x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +lean_ctor_set(x_21, 1, x_25); +x_26 = lean_st_ref_set(x_5, x_21, x_22); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +x_29 = lean_box(0); +lean_ctor_set(x_26, 0, x_29); +return x_26; } -lean_ctor_set(x_373, 0, x_372); -lean_ctor_set(x_373, 1, x_364); -x_374 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_374, 0, x_267); -lean_ctor_set(x_374, 1, x_373); -if (lean_is_scalar(x_353)) { - x_375 = lean_alloc_ctor(1, 1, 0); -} else { - x_375 = x_353; +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_26, 1); +lean_inc(x_30); +lean_dec(x_26); +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +return x_32; } -lean_ctor_set(x_375, 0, x_374); -x_18 = x_375; -x_19 = x_275; -goto block_24; } else { -size_t x_376; size_t x_377; lean_object* x_378; uint8_t x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; -x_376 = 0; -x_377 = lean_usize_of_nat(x_360); -lean_dec(x_360); -x_378 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_4, x_359, x_376, x_377, x_267); -lean_dec(x_359); -x_379 = 1; -x_380 = lean_box(x_379); -if (lean_is_scalar(x_270)) { - x_381 = lean_alloc_ctor(0, 2, 0); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 2); +x_35 = lean_ctor_get(x_21, 3); +x_36 = lean_ctor_get(x_21, 4); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_21); +x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_34); +lean_ctor_set(x_38, 3, x_35); +lean_ctor_set(x_38, 4, x_36); +x_39 = lean_st_ref_set(x_5, x_38, x_22); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; } else { - x_381 = x_270; + lean_dec_ref(x_39); + x_41 = lean_box(0); } -lean_ctor_set(x_381, 0, x_380); -lean_ctor_set(x_381, 1, x_364); -x_382 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_382, 0, x_378); -lean_ctor_set(x_382, 1, x_381); -if (lean_is_scalar(x_353)) { - x_383 = lean_alloc_ctor(1, 1, 0); +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); } else { - x_383 = x_353; -} -lean_ctor_set(x_383, 0, x_382); -x_18 = x_383; -x_19 = x_275; -goto block_24; -} + x_43 = x_41; } +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } else { -lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; -lean_dec(x_271); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +x_46 = lean_ctor_get(x_11, 2); +x_47 = lean_ctor_get(x_11, 3); +x_48 = lean_ctor_get(x_11, 5); +x_49 = lean_ctor_get(x_11, 6); +x_50 = lean_ctor_get(x_11, 7); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); -x_384 = lean_ctor_get(x_273, 0); -lean_inc(x_384); -x_385 = lean_ctor_get(x_273, 1); -lean_inc(x_385); -if (lean_is_exclusive(x_273)) { - lean_ctor_release(x_273, 0); - lean_ctor_release(x_273, 1); - x_386 = x_273; +x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); +x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; +x_53 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_45); +lean_ctor_set(x_53, 2, x_46); +lean_ctor_set(x_53, 3, x_47); +lean_ctor_set(x_53, 4, x_52); +lean_ctor_set(x_53, 5, x_48); +lean_ctor_set(x_53, 6, x_49); +lean_ctor_set(x_53, 7, x_50); +x_54 = lean_st_ref_set(x_7, x_53, x_12); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_st_ref_take(x_5, x_55); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = lean_ctor_get(x_57, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_57, 2); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 3); +lean_inc(x_61); +x_62 = lean_ctor_get(x_57, 4); +lean_inc(x_62); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + lean_ctor_release(x_57, 2); + lean_ctor_release(x_57, 3); + lean_ctor_release(x_57, 4); + x_63 = x_57; } else { - lean_dec_ref(x_273); - x_386 = lean_box(0); + lean_dec_ref(x_57); + x_63 = lean_box(0); } -if (lean_is_scalar(x_386)) { - x_387 = lean_alloc_ctor(1, 2, 0); +x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 5, 0); } else { - x_387 = x_386; -} -lean_ctor_set(x_387, 0, x_384); -lean_ctor_set(x_387, 1, x_385); -return x_387; -} + x_65 = x_63; } -else -{ -lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; -lean_dec(x_271); -lean_dec(x_17); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - x_388 = x_272; +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_60); +lean_ctor_set(x_65, 3, x_61); +lean_ctor_set(x_65, 4, x_62); +x_66 = lean_st_ref_set(x_5, x_65, x_58); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_68 = x_66; } else { - lean_dec_ref(x_272); - x_388 = lean_box(0); + lean_dec_ref(x_66); + x_68 = lean_box(0); } -if (lean_is_scalar(x_270)) { - x_389 = lean_alloc_ctor(0, 2, 0); +x_69 = lean_box(0); +if (lean_is_scalar(x_68)) { + x_70 = lean_alloc_ctor(0, 2, 0); } else { - x_389 = x_270; + x_70 = x_68; } -lean_ctor_set(x_389, 0, x_268); -lean_ctor_set(x_389, 1, x_269); -x_390 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_390, 0, x_267); -lean_ctor_set(x_390, 1, x_389); -if (lean_is_scalar(x_388)) { - x_391 = lean_alloc_ctor(1, 1, 0); -} else { - x_391 = x_388; +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_67); +return x_70; } -lean_ctor_set(x_391, 0, x_390); -x_18 = x_391; -x_19 = x_14; -goto block_24; } } -block_24: +static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = l_Lean_Meta_Grind_mkEMatchTheorem(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_13 = 0; +x_14 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_12, x_10, x_13, x_4, x_5, x_6, x_7, x_11); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_14; +} +else +{ +uint8_t x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +return x_9; +} +else { -lean_object* x_20; size_t x_21; size_t x_22; -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = 1; -x_22 = lean_usize_add(x_8, x_21); -x_8 = x_22; -x_9 = x_20; -x_14 = x_19; -goto _start; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_14; -x_14 = lean_usize_dec_lt(x_7, x_6); -if (x_14 == 0) +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); +lean_dec(x_3); +x_10 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchEqTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_15; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +uint8_t x_7; lean_object* x_8; +x_7 = 1; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_8 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_1, x_7, x_7, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_12 = 0; +x_13 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_11, x_9, x_12, x_2, x_3, x_4, x_5, x_10); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -return x_15; +return x_13; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_24; -x_16 = lean_array_uget(x_5, x_7); -x_24 = !lean_is_exclusive(x_8); -if (x_24 == 0) +uint8_t x_14; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_14 = !lean_is_exclusive(x_8); +if (x_14 == 0) { -lean_object* x_25; uint8_t x_26; -x_25 = lean_ctor_get(x_8, 1); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +return x_8; +} +else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_8, 0); -x_28 = lean_ctor_get(x_25, 0); -x_29 = lean_ctor_get(x_25, 1); -x_30 = l_Lean_Expr_fvarId_x21(x_16); -x_31 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_29, x_30); -if (lean_obj_tag(x_31) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_8, 0); +x_16 = lean_ctor_get(x_8, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_8); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_32; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_32 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_32) == 0) +lean_object* x_3; uint8_t x_4; +x_3 = lean_st_ref_get(x_1, x_2); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -x_35 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_27, x_30); -if (lean_obj_tag(x_35) == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_8 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_9 = l_Lean_ScopedEnvExtension_getState___rarg(x_7, x_8, x_6); +lean_dec(x_6); +lean_ctor_set(x_3, 0, x_9); +return x_3; +} +else { -lean_object* x_36; -lean_inc(x_12); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_33); -x_36 = l_Lean_Meta_isProp(x_33, x_9, x_10, x_11, x_12, x_34); -if (lean_obj_tag(x_36) == 0) -{ -lean_object* x_37; uint8_t x_38; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_unbox(x_37); -lean_dec(x_37); -if (x_38 == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_dec(x_36); -lean_inc(x_9); -lean_inc(x_30); -x_40 = l_Lean_FVarId_getDecl(x_30, x_9, x_10, x_11, x_12, x_39); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_LocalDecl_binderInfo(x_41); -lean_dec(x_41); -x_44 = lean_box(x_43); -if (lean_obj_tag(x_44) == 3) -{ -lean_object* x_45; +lean_dec(x_3); +x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_33); -lean_inc(x_27); -lean_inc(x_3); -x_45 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_27, x_33, x_9, x_10, x_11, x_12, x_42); -if (lean_obj_tag(x_45) == 0) +lean_dec(x_10); +x_13 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_14 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_15 = l_Lean_ScopedEnvExtension_getState___rarg(x_13, x_14, x_12); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_11); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems(lean_object* x_1) { +_start: { -lean_object* x_46; uint8_t x_47; -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_unbox(x_46); -lean_dec(x_46); -if (x_47 == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_48; lean_object* x_49; -lean_dec(x_33); -lean_dec(x_30); -x_48 = lean_ctor_get(x_45, 1); -lean_inc(x_48); -lean_dec(x_45); -x_49 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_49, 0, x_8); -x_17 = x_49; -x_18 = x_48; -goto block_23; +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_getEMatchTheorems___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___boxed(lean_object* x_1) { +_start: { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; -lean_dec(x_28); -x_50 = lean_ctor_get(x_45, 1); -lean_inc(x_50); -lean_dec(x_45); -x_51 = lean_box(0); -lean_inc(x_30); -x_52 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_51); -x_53 = lean_box(0); -lean_inc(x_2); -x_54 = lean_array_mk(x_2); -x_55 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_56 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -lean_ctor_set(x_56, 2, x_54); -x_57 = l_Lean_CollectFVars_main(x_33, x_56); -x_58 = lean_ctor_get(x_57, 2); -lean_inc(x_58); -lean_dec(x_57); -x_59 = lean_array_get_size(x_58); -x_60 = lean_unsigned_to_nat(0u); -x_61 = lean_nat_dec_lt(x_60, x_59); -x_62 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_51); -if (x_61 == 0) +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_getEMatchTheorems(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: { -uint8_t x_63; lean_object* x_64; lean_object* x_65; -lean_dec(x_59); -lean_dec(x_58); -x_63 = 1; -x_64 = lean_box(x_63); -lean_ctor_set(x_25, 1, x_62); -lean_ctor_set(x_25, 0, x_64); -lean_ctor_set(x_8, 0, x_52); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_8); -x_17 = x_65; -x_18 = x_50; -goto block_23; +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; } -else +case 2: { -uint8_t x_66; -x_66 = lean_nat_dec_le(x_59, x_59); -if (x_66 == 0) +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +case 3: { -uint8_t x_67; lean_object* x_68; lean_object* x_69; -lean_dec(x_59); -lean_dec(x_58); -x_67 = 1; -x_68 = lean_box(x_67); -lean_ctor_set(x_25, 1, x_62); -lean_ctor_set(x_25, 0, x_68); -lean_ctor_set(x_8, 0, x_52); -x_69 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_69, 0, x_8); -x_17 = x_69; -x_18 = x_50; -goto block_23; +lean_object* x_5; +x_5 = lean_unsigned_to_nat(3u); +return x_5; } -else +case 4: { -size_t x_70; size_t x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; -x_70 = 0; -x_71 = lean_usize_of_nat(x_59); -lean_dec(x_59); -x_72 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_58, x_70, x_71, x_52); -lean_dec(x_58); -x_73 = 1; -x_74 = lean_box(x_73); -lean_ctor_set(x_25, 1, x_62); -lean_ctor_set(x_25, 0, x_74); -lean_ctor_set(x_8, 0, x_72); -x_75 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_75, 0, x_8); -x_17 = x_75; -x_18 = x_50; -goto block_23; +lean_object* x_6; +x_6 = lean_unsigned_to_nat(4u); +return x_6; +} +default: +{ +lean_object* x_7; +x_7 = lean_unsigned_to_nat(5u); +return x_7; } } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx___boxed(lean_object* x_1) { +_start: { -uint8_t x_76; -lean_dec(x_33); -lean_dec(x_30); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_28); -lean_free_object(x_8); -lean_dec(x_27); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_76 = !lean_is_exclusive(x_45); -if (x_76 == 0) +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: { -return x_45; +lean_inc(x_1); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1() { +_start: { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_45, 0); -x_78 = lean_ctor_get(x_45, 1); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_45); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1; +return x_4; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion(lean_object* x_1) { +_start: { -lean_object* x_80; -lean_dec(x_44); -lean_dec(x_33); -lean_dec(x_30); -x_80 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_80, 0, x_8); -x_17 = x_80; -x_18 = x_42; -goto block_23; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___boxed), 3, 0); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: { -uint8_t x_81; -lean_dec(x_33); -lean_dec(x_30); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_28); -lean_free_object(x_8); -lean_dec(x_27); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); lean_dec(x_2); -x_81 = !lean_is_exclusive(x_40); -if (x_81 == 0) +x_6 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static uint8_t _init_l_Lean_Meta_Grind_instInhabitedTheoremKind() { +_start: { -return x_40; +uint8_t x_1; +x_1 = 0; +return x_1; } -else +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(uint8_t x_1, uint8_t x_2) { +_start: { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_40, 0); -x_83 = lean_ctor_get(x_40, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_40); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -return x_84; +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx(x_1); +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toCtorIdx(x_2); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = lean_unbox(x_2); +lean_dec(x_2); +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; } } -else +static lean_object* _init_l_Lean_Meta_Grind_instBEqTheoremKind___closed__1() { +_start: { -lean_object* x_85; uint8_t x_86; -x_85 = lean_ctor_get(x_36, 1); -lean_inc(x_85); -lean_dec(x_36); -x_86 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_27, x_33); -if (x_86 == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instBEqTheoremKind() { +_start: { -lean_object* x_87; -lean_dec(x_30); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_8); -x_17 = x_87; -x_18 = x_85; -goto block_23; +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instBEqTheoremKind___closed__1; +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1() { +_start: { -lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_28); -x_88 = lean_box(0); -lean_inc(x_30); -x_89 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_27, x_30, x_88); -x_90 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_88); -x_91 = 1; -x_92 = lean_box(x_91); -lean_ctor_set(x_25, 1, x_90); -lean_ctor_set(x_25, 0, x_92); -lean_ctor_set(x_8, 0, x_89); -x_93 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_93, 0, x_8); -x_17 = x_93; -x_18 = x_85; -goto block_23; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind =]", 9, 9); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind =_]", 10, 10); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3() { +_start: { -uint8_t x_94; -lean_dec(x_33); -lean_dec(x_30); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_28); -lean_free_object(x_8); -lean_dec(x_27); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_94 = !lean_is_exclusive(x_36); -if (x_94 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind _=_]", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4() { +_start: { -return x_36; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind →]", 11, 9); +return x_1; +} } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5() { +_start: { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_36, 0); -x_96 = lean_ctor_get(x_36, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_36); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind ←]", 11, 9); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[grind]", 7, 7); +return x_1; +} } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(uint8_t x_1) { +_start: { -uint8_t x_98; -lean_dec(x_28); -x_98 = !lean_is_exclusive(x_35); -if (x_98 == 0) +switch (x_1) { +case 0: { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; -x_99 = lean_ctor_get(x_35, 0); -lean_dec(x_99); -x_100 = lean_box(0); -lean_inc(x_2); -x_101 = lean_array_mk(x_2); -x_102 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_103 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_100); -lean_ctor_set(x_103, 2, x_101); -x_104 = l_Lean_CollectFVars_main(x_33, x_103); -x_105 = lean_ctor_get(x_104, 2); -lean_inc(x_105); -lean_dec(x_104); -x_106 = lean_array_get_size(x_105); -x_107 = lean_unsigned_to_nat(0u); -x_108 = lean_nat_dec_lt(x_107, x_106); -x_109 = lean_box(0); -x_110 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_109); -if (x_108 == 0) +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1; +return x_2; +} +case 1: { -uint8_t x_111; lean_object* x_112; -lean_dec(x_106); -lean_dec(x_105); -x_111 = 1; -x_112 = lean_box(x_111); -lean_ctor_set(x_25, 1, x_110); -lean_ctor_set(x_25, 0, x_112); -lean_ctor_set(x_35, 0, x_8); -x_17 = x_35; -x_18 = x_34; -goto block_23; +lean_object* x_3; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2; +return x_3; } -else +case 2: { -uint8_t x_113; -x_113 = lean_nat_dec_le(x_106, x_106); -if (x_113 == 0) +lean_object* x_4; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3; +return x_4; +} +case 3: { -uint8_t x_114; lean_object* x_115; -lean_dec(x_106); -lean_dec(x_105); -x_114 = 1; -x_115 = lean_box(x_114); -lean_ctor_set(x_25, 1, x_110); -lean_ctor_set(x_25, 0, x_115); -lean_ctor_set(x_35, 0, x_8); -x_17 = x_35; -x_18 = x_34; -goto block_23; +lean_object* x_5; +x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4; +return x_5; } -else +case 4: { -size_t x_116; size_t x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; -x_116 = 0; -x_117 = lean_usize_of_nat(x_106); -lean_dec(x_106); -x_118 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_105, x_116, x_117, x_27); -lean_dec(x_105); -x_119 = 1; -x_120 = lean_box(x_119); -lean_ctor_set(x_25, 1, x_110); -lean_ctor_set(x_25, 0, x_120); -lean_ctor_set(x_8, 0, x_118); -lean_ctor_set(x_35, 0, x_8); -x_17 = x_35; -x_18 = x_34; -goto block_23; +lean_object* x_6; +x_6 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5; +return x_6; } +default: +{ +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6; +return x_7; } } -else -{ -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; -lean_dec(x_35); -x_121 = lean_box(0); -lean_inc(x_2); -x_122 = lean_array_mk(x_2); -x_123 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_124 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_121); -lean_ctor_set(x_124, 2, x_122); -x_125 = l_Lean_CollectFVars_main(x_33, x_124); -x_126 = lean_ctor_get(x_125, 2); -lean_inc(x_126); -lean_dec(x_125); -x_127 = lean_array_get_size(x_126); -x_128 = lean_unsigned_to_nat(0u); -x_129 = lean_nat_dec_lt(x_128, x_127); -x_130 = lean_box(0); -x_131 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_29, x_30, x_130); -if (x_129 == 0) +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___boxed(lean_object* x_1) { +_start: { -uint8_t x_132; lean_object* x_133; lean_object* x_134; -lean_dec(x_127); -lean_dec(x_126); -x_132 = 1; -x_133 = lean_box(x_132); -lean_ctor_set(x_25, 1, x_131); -lean_ctor_set(x_25, 0, x_133); -x_134 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_134, 0, x_8); -x_17 = x_134; -x_18 = x_34; -goto block_23; +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1() { +_start: { -uint8_t x_135; -x_135 = lean_nat_dec_le(x_127, x_127); -if (x_135 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find pattern in the left-hand side of the theorem's conclusion", 72, 72); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2() { +_start: { -uint8_t x_136; lean_object* x_137; lean_object* x_138; -lean_dec(x_127); -lean_dec(x_126); -x_136 = 1; -x_137 = lean_box(x_136); -lean_ctor_set(x_25, 1, x_131); -lean_ctor_set(x_25, 0, x_137); -x_138 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_138, 0, x_8); -x_17 = x_138; -x_18 = x_34; -goto block_23; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find pattern in the right-hand side of the theorem's conclusion", 73, 73); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3() { +_start: { -size_t x_139; size_t x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; -x_139 = 0; -x_140 = lean_usize_of_nat(x_127); -lean_dec(x_127); -x_141 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_126, x_139, x_140, x_27); -lean_dec(x_126); -x_142 = 1; -x_143 = lean_box(x_142); -lean_ctor_set(x_25, 1, x_131); -lean_ctor_set(x_25, 0, x_143); -lean_ctor_set(x_8, 0, x_141); -x_144 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_144, 0, x_8); -x_17 = x_144; -x_18 = x_34; -goto block_23; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.TheoremKind.explainFailure", 90, 90); +return x_1; +} } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3; +x_3 = lean_unsigned_to_nat(516u); +x_4 = lean_unsigned_to_nat(16u); +x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find patterns in the antecedents of the theorem", 57, 57); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6() { +_start: { -uint8_t x_145; -lean_dec(x_30); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_28); -lean_free_object(x_8); -lean_dec(x_27); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_145 = !lean_is_exclusive(x_32); -if (x_145 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find patterns in the theorem's conclusion", 51, 51); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7() { +_start: { -return x_32; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("failed to find patterns", 23, 23); +return x_1; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure(uint8_t x_1) { +_start: { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_32, 0); -x_147 = lean_ctor_get(x_32, 1); -lean_inc(x_147); -lean_inc(x_146); -lean_dec(x_32); -x_148 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_147); -return x_148; +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1; +return x_2; } +case 1: +{ +lean_object* x_3; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2; +return x_3; } +case 2: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4; +x_5 = l_panic___at_String_fromUTF8_x21___spec__1(x_4); +return x_5; } -else +case 3: { -uint8_t x_149; -lean_dec(x_30); -lean_dec(x_16); -x_149 = !lean_is_exclusive(x_31); -if (x_149 == 0) +lean_object* x_6; +x_6 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5; +return x_6; +} +case 4: { -lean_object* x_150; -x_150 = lean_ctor_get(x_31, 0); -lean_dec(x_150); -lean_ctor_set(x_31, 0, x_8); -x_17 = x_31; -x_18 = x_13; -goto block_23; +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6; +return x_7; } -else +default: { -lean_object* x_151; -lean_dec(x_31); -x_151 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_151, 0, x_8); -x_17 = x_151; -x_18 = x_13; -goto block_23; +lean_object* x_8; +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7; +return x_8; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___boxed(lean_object* x_1) { +_start: { -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_152 = lean_ctor_get(x_8, 0); -x_153 = lean_ctor_get(x_25, 0); -x_154 = lean_ctor_get(x_25, 1); -lean_inc(x_154); -lean_inc(x_153); -lean_dec(x_25); -x_155 = l_Lean_Expr_fvarId_x21(x_16); -x_156 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_154, x_155); -if (lean_obj_tag(x_156) == 0) +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_157; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_157 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_157) == 0) +uint8_t x_10; +x_10 = lean_usize_dec_eq(x_2, x_3); +if (x_10 == 0) { -lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_158 = lean_ctor_get(x_157, 0); -lean_inc(x_158); -x_159 = lean_ctor_get(x_157, 1); -lean_inc(x_159); -lean_dec(x_157); -x_160 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_152, x_155); -if (lean_obj_tag(x_160) == 0) +lean_object* x_11; lean_object* x_12; +x_11 = lean_array_uget(x_1, x_2); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_12 = lean_infer_type(x_11, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_161; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_158); -x_161 = l_Lean_Meta_isProp(x_158, x_9, x_10, x_11, x_12, x_159); -if (lean_obj_tag(x_161) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_13); +x_15 = l_Lean_Meta_isProp(x_13, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_162; uint8_t x_163; -x_162 = lean_ctor_get(x_161, 0); -lean_inc(x_162); -x_163 = lean_unbox(x_162); -lean_dec(x_162); -if (x_163 == 0) +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) { -lean_object* x_164; lean_object* x_165; -x_164 = lean_ctor_get(x_161, 1); -lean_inc(x_164); -lean_dec(x_161); -lean_inc(x_9); -lean_inc(x_155); -x_165 = l_Lean_FVarId_getDecl(x_155, x_9, x_10, x_11, x_12, x_164); -if (lean_obj_tag(x_165) == 0) +lean_object* x_18; size_t x_19; size_t x_20; +lean_dec(x_13); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = 1; +x_20 = lean_usize_add(x_2, x_19); +x_2 = x_20; +x_9 = x_18; +goto _start; +} +else { -lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; -x_166 = lean_ctor_get(x_165, 0); -lean_inc(x_166); -x_167 = lean_ctor_get(x_165, 1); -lean_inc(x_167); -lean_dec(x_165); -x_168 = l_Lean_LocalDecl_binderInfo(x_166); -lean_dec(x_166); -x_169 = lean_box(x_168); -if (lean_obj_tag(x_169) == 3) +lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_dec(x_15); +x_23 = lean_array_push(x_4, x_13); +x_24 = 1; +x_25 = lean_usize_add(x_2, x_24); +x_2 = x_25; +x_4 = x_23; +x_9 = x_22; +goto _start; +} +} +else { -lean_object* x_170; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_158); -lean_inc(x_152); -lean_inc(x_3); -x_170 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_152, x_158, x_9, x_10, x_11, x_12, x_167); -if (lean_obj_tag(x_170) == 0) +uint8_t x_27; +lean_dec(x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_27 = !lean_is_exclusive(x_15); +if (x_27 == 0) { -lean_object* x_171; uint8_t x_172; -x_171 = lean_ctor_get(x_170, 0); -lean_inc(x_171); -x_172 = lean_unbox(x_171); -lean_dec(x_171); -if (x_172 == 0) +return x_15; +} +else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -lean_dec(x_158); -lean_dec(x_155); -x_173 = lean_ctor_get(x_170, 1); -lean_inc(x_173); -lean_dec(x_170); -x_174 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_174, 0, x_153); -lean_ctor_set(x_174, 1, x_154); -lean_ctor_set(x_8, 1, x_174); -x_175 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_175, 0, x_8); -x_17 = x_175; -x_18 = x_173; -goto block_23; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_15, 0); +x_29 = lean_ctor_get(x_15, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_15); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} } else { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; -lean_dec(x_153); -x_176 = lean_ctor_get(x_170, 1); -lean_inc(x_176); -lean_dec(x_170); -x_177 = lean_box(0); -lean_inc(x_155); -x_178 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_177); -x_179 = lean_box(0); -lean_inc(x_2); -x_180 = lean_array_mk(x_2); -x_181 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_182 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_182, 0, x_181); -lean_ctor_set(x_182, 1, x_179); -lean_ctor_set(x_182, 2, x_180); -x_183 = l_Lean_CollectFVars_main(x_158, x_182); -x_184 = lean_ctor_get(x_183, 2); -lean_inc(x_184); -lean_dec(x_183); -x_185 = lean_array_get_size(x_184); -x_186 = lean_unsigned_to_nat(0u); -x_187 = lean_nat_dec_lt(x_186, x_185); -x_188 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_177); -if (x_187 == 0) +uint8_t x_31; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_31 = !lean_is_exclusive(x_12); +if (x_31 == 0) { -uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -lean_dec(x_185); -lean_dec(x_184); -x_189 = 1; -x_190 = lean_box(x_189); -x_191 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_191, 0, x_190); -lean_ctor_set(x_191, 1, x_188); -lean_ctor_set(x_8, 1, x_191); -lean_ctor_set(x_8, 0, x_178); -x_192 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_192, 0, x_8); -x_17 = x_192; -x_18 = x_176; -goto block_23; +return x_12; } else { -uint8_t x_193; -x_193 = lean_nat_dec_le(x_185, x_185); -if (x_193 == 0) -{ -uint8_t x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; -lean_dec(x_185); -lean_dec(x_184); -x_194 = 1; -x_195 = lean_box(x_194); -x_196 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_196, 0, x_195); -lean_ctor_set(x_196, 1, x_188); -lean_ctor_set(x_8, 1, x_196); -lean_ctor_set(x_8, 0, x_178); -x_197 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_197, 0, x_8); -x_17 = x_197; -x_18 = x_176; -goto block_23; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_12, 0); +x_33 = lean_ctor_get(x_12, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_12); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} } else { -size_t x_198; size_t x_199; lean_object* x_200; uint8_t x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_198 = 0; -x_199 = lean_usize_of_nat(x_185); -lean_dec(x_185); -x_200 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_184, x_198, x_199, x_178); -lean_dec(x_184); -x_201 = 1; -x_202 = lean_box(x_201); -x_203 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_203, 0, x_202); -lean_ctor_set(x_203, 1, x_188); -lean_ctor_set(x_8, 1, x_203); -lean_ctor_set(x_8, 0, x_200); -x_204 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_204, 0, x_8); -x_17 = x_204; -x_18 = x_176; -goto block_23; +lean_object* x_35; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_4); +lean_ctor_set(x_35, 1, x_9); +return x_35; } } } +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = lean_nat_dec_lt(x_2, x_3); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_10 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_8); +return x_11; } else { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; -lean_dec(x_158); -lean_dec(x_155); -lean_dec(x_154); -lean_dec(x_153); -lean_free_object(x_8); -lean_dec(x_152); +lean_object* x_12; uint8_t x_13; +x_12 = lean_array_get_size(x_1); +x_13 = lean_nat_dec_le(x_3, x_12); lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_205 = lean_ctor_get(x_170, 0); -lean_inc(x_205); -x_206 = lean_ctor_get(x_170, 1); -lean_inc(x_206); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - lean_ctor_release(x_170, 1); - x_207 = x_170; -} else { - lean_dec_ref(x_170); - x_207 = lean_box(0); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_14 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_8); +return x_15; } -if (lean_is_scalar(x_207)) { - x_208 = lean_alloc_ctor(1, 2, 0); -} else { - x_208 = x_207; +else +{ +size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_usize_of_nat(x_2); +x_17 = lean_usize_of_nat(x_3); +x_18 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_19 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2(x_1, x_16, x_17, x_18, x_4, x_5, x_6, x_7, x_8); +return x_19; } -lean_ctor_set(x_208, 0, x_205); -lean_ctor_set(x_208, 1, x_206); -return x_208; } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_209; lean_object* x_210; -lean_dec(x_169); -lean_dec(x_158); -lean_dec(x_155); -x_209 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_209, 0, x_153); -lean_ctor_set(x_209, 1, x_154); -lean_ctor_set(x_8, 1, x_209); -x_210 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_210, 0, x_8); -x_17 = x_210; -x_18 = x_167; -goto block_23; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_array_get_size(x_1); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1(x_1, x_8, x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_7); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -lean_dec(x_158); -lean_dec(x_155); -lean_dec(x_154); -lean_dec(x_153); -lean_free_object(x_8); -lean_dec(x_152); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_11 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_12 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__2(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Array_filterMapM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); lean_dec(x_2); -x_211 = lean_ctor_get(x_165, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_165, 1); -lean_inc(x_212); -if (lean_is_exclusive(x_165)) { - lean_ctor_release(x_165, 0); - lean_ctor_release(x_165, 1); - x_213 = x_165; -} else { - lean_dec_ref(x_165); - x_213 = lean_box(0); +lean_dec(x_1); +return x_9; } -if (lean_is_scalar(x_213)) { - x_214 = lean_alloc_ctor(1, 2, 0); -} else { - x_214 = x_213; } -lean_ctor_set(x_214, 0, x_211); -lean_ctor_set(x_214, 1, x_212); -return x_214; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_1); +return x_7; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_215; uint8_t x_216; -x_215 = lean_ctor_get(x_161, 1); -lean_inc(x_215); -lean_dec(x_161); -x_216 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_152, x_158); -if (x_216 == 0) +switch (lean_obj_tag(x_1)) { +case 1: { -lean_object* x_217; lean_object* x_218; -lean_dec(x_155); -x_217 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_217, 0, x_153); -lean_ctor_set(x_217, 1, x_154); -lean_ctor_set(x_8, 1, x_217); -x_218 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_218, 0, x_8); -x_17 = x_218; -x_18 = x_215; -goto block_23; +lean_object* x_10; uint8_t x_11; +x_10 = lean_ctor_get(x_2, 1); +x_11 = l_Array_contains___at_Lean_Meta_arrowDomainsN___spec__2(x_10, x_1); +if (x_11 == 0) +{ +uint8_t x_12; lean_object* x_13; lean_object* x_14; +x_12 = 1; +x_13 = lean_box(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_9); +return x_14; } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; uint8_t x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; -lean_dec(x_153); -x_219 = lean_box(0); -lean_inc(x_155); -x_220 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_152, x_155, x_219); -x_221 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_219); -x_222 = 1; -x_223 = lean_box(x_222); -x_224 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_224, 0, x_223); -lean_ctor_set(x_224, 1, x_221); -lean_ctor_set(x_8, 1, x_224); -lean_ctor_set(x_8, 0, x_220); -x_225 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_225, 0, x_8); -x_17 = x_225; -x_18 = x_215; -goto block_23; +uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_15 = 0; +x_16 = lean_box(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_9); +return x_17; } } +case 4: +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_1, 0); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_forbiddenDeclNames; +x_20 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_19, x_18); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_21 = 1; +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_9); +return x_23; } else { -lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -lean_dec(x_158); -lean_dec(x_155); -lean_dec(x_154); -lean_dec(x_153); -lean_free_object(x_8); -lean_dec(x_152); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_226 = lean_ctor_get(x_161, 0); -lean_inc(x_226); -x_227 = lean_ctor_get(x_161, 1); -lean_inc(x_227); -if (lean_is_exclusive(x_161)) { - lean_ctor_release(x_161, 0); - lean_ctor_release(x_161, 1); - x_228 = x_161; -} else { - lean_dec_ref(x_161); - x_228 = lean_box(0); +uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_24 = 0; +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_9); +return x_26; +} +} +default: +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_9); +return x_29; +} } -if (lean_is_scalar(x_228)) { - x_229 = lean_alloc_ctor(1, 2, 0); -} else { - x_229 = x_228; } -lean_ctor_set(x_229, 0, x_226); -lean_ctor_set(x_229, 1, x_227); -return x_229; +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; } } +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_ctor_get(x_7, 2); +x_11 = l_Lean_isTracingEnabledForCore(x_1, x_10, x_9); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +return x_11; +} else { -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; lean_object* x_240; lean_object* x_241; -lean_dec(x_153); -if (lean_is_exclusive(x_160)) { - lean_ctor_release(x_160, 0); - x_230 = x_160; -} else { - lean_dec_ref(x_160); - x_230 = lean_box(0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; } -x_231 = lean_box(0); -lean_inc(x_2); -x_232 = lean_array_mk(x_2); -x_233 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_234 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_234, 0, x_233); -lean_ctor_set(x_234, 1, x_231); -lean_ctor_set(x_234, 2, x_232); -x_235 = l_Lean_CollectFVars_main(x_158, x_234); -x_236 = lean_ctor_get(x_235, 2); -lean_inc(x_236); -lean_dec(x_235); -x_237 = lean_array_get_size(x_236); -x_238 = lean_unsigned_to_nat(0u); -x_239 = lean_nat_dec_lt(x_238, x_237); -x_240 = lean_box(0); -x_241 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_154, x_155, x_240); -if (x_239 == 0) +} +} +static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1() { +_start: { -uint8_t x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; -lean_dec(x_237); -lean_dec(x_236); -x_242 = 1; -x_243 = lean_box(x_242); -x_244 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_244, 0, x_243); -lean_ctor_set(x_244, 1, x_241); -lean_ctor_set(x_8, 1, x_244); -if (lean_is_scalar(x_230)) { - x_245 = lean_alloc_ctor(1, 1, 0); -} else { - x_245 = x_230; +lean_object* x_1; uint8_t x_2; double x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = l_Float_ofScientific(x_1, x_2, x_1); +return x_3; } -lean_ctor_set(x_245, 0, x_8); -x_17 = x_245; -x_18 = x_159; -goto block_23; } -else +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_246; -x_246 = lean_nat_dec_le(x_237, x_237); -if (x_246 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_11 = lean_ctor_get(x_8, 5); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_6, x_7, x_8, x_9, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_st_ref_take(x_9, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_16, 3); +lean_inc(x_17); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -uint8_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; -lean_dec(x_237); -lean_dec(x_236); -x_247 = 1; -x_248 = lean_box(x_247); -x_249 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_249, 0, x_248); -lean_ctor_set(x_249, 1, x_241); -lean_ctor_set(x_8, 1, x_249); -if (lean_is_scalar(x_230)) { - x_250 = lean_alloc_ctor(1, 1, 0); -} else { - x_250 = x_230; +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 1); +x_20 = lean_ctor_get(x_15, 0); +lean_dec(x_20); +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_16, 3); +lean_dec(x_22); +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +lean_object* x_24; double x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_24 = lean_ctor_get(x_17, 0); +x_25 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +x_26 = 0; +x_27 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_28 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_28, 0, x_1); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set_float(x_28, sizeof(void*)*2, x_25); +lean_ctor_set_float(x_28, sizeof(void*)*2 + 8, x_25); +lean_ctor_set_uint8(x_28, sizeof(void*)*2 + 16, x_26); +x_29 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_30 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_13); +lean_ctor_set(x_30, 2, x_29); +lean_inc(x_11); +lean_ctor_set(x_15, 1, x_30); +lean_ctor_set(x_15, 0, x_11); +x_31 = l_Lean_PersistentArray_push___rarg(x_24, x_15); +lean_ctor_set(x_17, 0, x_31); +x_32 = lean_st_ref_set(x_9, x_16, x_19); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +x_35 = lean_box(0); +lean_ctor_set(x_32, 0, x_35); +return x_32; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_32, 1); +lean_inc(x_36); +lean_dec(x_32); +x_37 = lean_box(0); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; } -lean_ctor_set(x_250, 0, x_8); -x_17 = x_250; -x_18 = x_159; -goto block_23; } else { -size_t x_251; size_t x_252; lean_object* x_253; uint8_t x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; -x_251 = 0; -x_252 = lean_usize_of_nat(x_237); -lean_dec(x_237); -x_253 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_236, x_251, x_252, x_152); -lean_dec(x_236); -x_254 = 1; -x_255 = lean_box(x_254); -x_256 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_256, 0, x_255); -lean_ctor_set(x_256, 1, x_241); -lean_ctor_set(x_8, 1, x_256); -lean_ctor_set(x_8, 0, x_253); -if (lean_is_scalar(x_230)) { - x_257 = lean_alloc_ctor(1, 1, 0); +uint64_t x_39; lean_object* x_40; double x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_39 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_40 = lean_ctor_get(x_17, 0); +lean_inc(x_40); +lean_dec(x_17); +x_41 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +x_42 = 0; +x_43 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_44 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_44, 0, x_1); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set_float(x_44, sizeof(void*)*2, x_41); +lean_ctor_set_float(x_44, sizeof(void*)*2 + 8, x_41); +lean_ctor_set_uint8(x_44, sizeof(void*)*2 + 16, x_42); +x_45 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_46 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_13); +lean_ctor_set(x_46, 2, x_45); +lean_inc(x_11); +lean_ctor_set(x_15, 1, x_46); +lean_ctor_set(x_15, 0, x_11); +x_47 = l_Lean_PersistentArray_push___rarg(x_40, x_15); +x_48 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set_uint64(x_48, sizeof(void*)*1, x_39); +lean_ctor_set(x_16, 3, x_48); +x_49 = lean_st_ref_set(x_9, x_16, x_19); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_51 = x_49; } else { - x_257 = x_230; -} -lean_ctor_set(x_257, 0, x_8); -x_17 = x_257; -x_18 = x_159; -goto block_23; + lean_dec_ref(x_49); + x_51 = lean_box(0); } +x_52 = lean_box(0); +if (lean_is_scalar(x_51)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_51; } +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_50); +return x_53; } } else { -lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; -lean_dec(x_155); -lean_dec(x_154); -lean_dec(x_153); -lean_free_object(x_8); -lean_dec(x_152); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_258 = lean_ctor_get(x_157, 0); -lean_inc(x_258); -x_259 = lean_ctor_get(x_157, 1); -lean_inc(x_259); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_260 = x_157; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint64_t x_61; lean_object* x_62; lean_object* x_63; double x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_54 = lean_ctor_get(x_16, 0); +x_55 = lean_ctor_get(x_16, 1); +x_56 = lean_ctor_get(x_16, 2); +x_57 = lean_ctor_get(x_16, 4); +x_58 = lean_ctor_get(x_16, 5); +x_59 = lean_ctor_get(x_16, 6); +x_60 = lean_ctor_get(x_16, 7); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_16); +x_61 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_62 = lean_ctor_get(x_17, 0); +lean_inc(x_62); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + x_63 = x_17; } else { - lean_dec_ref(x_157); - x_260 = lean_box(0); + lean_dec_ref(x_17); + x_63 = lean_box(0); } -if (lean_is_scalar(x_260)) { - x_261 = lean_alloc_ctor(1, 2, 0); +x_64 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +x_65 = 0; +x_66 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_67 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_67, 0, x_1); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set_float(x_67, sizeof(void*)*2, x_64); +lean_ctor_set_float(x_67, sizeof(void*)*2 + 8, x_64); +lean_ctor_set_uint8(x_67, sizeof(void*)*2 + 16, x_65); +x_68 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_69 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_13); +lean_ctor_set(x_69, 2, x_68); +lean_inc(x_11); +lean_ctor_set(x_15, 1, x_69); +lean_ctor_set(x_15, 0, x_11); +x_70 = l_Lean_PersistentArray_push___rarg(x_62, x_15); +if (lean_is_scalar(x_63)) { + x_71 = lean_alloc_ctor(0, 1, 8); } else { - x_261 = x_260; + x_71 = x_63; } -lean_ctor_set(x_261, 0, x_258); -lean_ctor_set(x_261, 1, x_259); -return x_261; +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set_uint64(x_71, sizeof(void*)*1, x_61); +x_72 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_72, 0, x_54); +lean_ctor_set(x_72, 1, x_55); +lean_ctor_set(x_72, 2, x_56); +lean_ctor_set(x_72, 3, x_71); +lean_ctor_set(x_72, 4, x_57); +lean_ctor_set(x_72, 5, x_58); +lean_ctor_set(x_72, 6, x_59); +lean_ctor_set(x_72, 7, x_60); +x_73 = lean_st_ref_set(x_9, x_72, x_19); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_75 = x_73; +} else { + lean_dec_ref(x_73); + x_75 = lean_box(0); +} +x_76 = lean_box(0); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); +} else { + x_77 = x_75; +} +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; -lean_dec(x_155); -lean_dec(x_16); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - x_262 = x_156; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint64_t x_87; lean_object* x_88; lean_object* x_89; double x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_78 = lean_ctor_get(x_15, 1); +lean_inc(x_78); +lean_dec(x_15); +x_79 = lean_ctor_get(x_16, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_16, 1); +lean_inc(x_80); +x_81 = lean_ctor_get(x_16, 2); +lean_inc(x_81); +x_82 = lean_ctor_get(x_16, 4); +lean_inc(x_82); +x_83 = lean_ctor_get(x_16, 5); +lean_inc(x_83); +x_84 = lean_ctor_get(x_16, 6); +lean_inc(x_84); +x_85 = lean_ctor_get(x_16, 7); +lean_inc(x_85); +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + lean_ctor_release(x_16, 1); + lean_ctor_release(x_16, 2); + lean_ctor_release(x_16, 3); + lean_ctor_release(x_16, 4); + lean_ctor_release(x_16, 5); + lean_ctor_release(x_16, 6); + lean_ctor_release(x_16, 7); + x_86 = x_16; } else { - lean_dec_ref(x_156); - x_262 = lean_box(0); + lean_dec_ref(x_16); + x_86 = lean_box(0); } -x_263 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_263, 0, x_153); -lean_ctor_set(x_263, 1, x_154); -lean_ctor_set(x_8, 1, x_263); -if (lean_is_scalar(x_262)) { - x_264 = lean_alloc_ctor(1, 1, 0); +x_87 = lean_ctor_get_uint64(x_17, sizeof(void*)*1); +x_88 = lean_ctor_get(x_17, 0); +lean_inc(x_88); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + x_89 = x_17; +} else { + lean_dec_ref(x_17); + x_89 = lean_box(0); +} +x_90 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1; +x_91 = 0; +x_92 = l_Lean_Meta_Grind_ppPattern___closed__3; +x_93 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_93, 0, x_1); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set_float(x_93, sizeof(void*)*2, x_90); +lean_ctor_set_float(x_93, sizeof(void*)*2 + 8, x_90); +lean_ctor_set_uint8(x_93, sizeof(void*)*2 + 16, x_91); +x_94 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_95 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_13); +lean_ctor_set(x_95, 2, x_94); +lean_inc(x_11); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_11); +lean_ctor_set(x_96, 1, x_95); +x_97 = l_Lean_PersistentArray_push___rarg(x_88, x_96); +if (lean_is_scalar(x_89)) { + x_98 = lean_alloc_ctor(0, 1, 8); } else { - x_264 = x_262; + x_98 = x_89; } -lean_ctor_set(x_264, 0, x_8); -x_17 = x_264; -x_18 = x_13; -goto block_23; +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set_uint64(x_98, sizeof(void*)*1, x_87); +if (lean_is_scalar(x_86)) { + x_99 = lean_alloc_ctor(0, 8, 0); +} else { + x_99 = x_86; +} +lean_ctor_set(x_99, 0, x_79); +lean_ctor_set(x_99, 1, x_80); +lean_ctor_set(x_99, 2, x_81); +lean_ctor_set(x_99, 3, x_98); +lean_ctor_set(x_99, 4, x_82); +lean_ctor_set(x_99, 5, x_83); +lean_ctor_set(x_99, 6, x_84); +lean_ctor_set(x_99, 7, x_85); +x_100 = lean_st_ref_set(x_9, x_99, x_78); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_102 = x_100; +} else { + lean_dec_ref(x_100); + x_102 = lean_box(0); } +x_103 = lean_box(0); +if (lean_is_scalar(x_102)) { + x_104 = lean_alloc_ctor(0, 2, 0); +} else { + x_104 = x_102; } +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_101); +return x_104; } -else -{ -lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -x_265 = lean_ctor_get(x_8, 1); -x_266 = lean_ctor_get(x_8, 0); -lean_inc(x_265); -lean_inc(x_266); -lean_dec(x_8); -x_267 = lean_ctor_get(x_265, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_265, 1); -lean_inc(x_268); -if (lean_is_exclusive(x_265)) { - lean_ctor_release(x_265, 0); - lean_ctor_release(x_265, 1); - x_269 = x_265; -} else { - lean_dec_ref(x_265); - x_269 = lean_box(0); } -x_270 = l_Lean_Expr_fvarId_x21(x_16); -x_271 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_268, x_270); -if (lean_obj_tag(x_271) == 0) -{ -lean_object* x_272; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_272 = lean_infer_type(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_272) == 0) +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_273; lean_object* x_274; lean_object* x_275; -x_273 = lean_ctor_get(x_272, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_272, 1); -lean_inc(x_274); -lean_dec(x_272); -x_275 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_266, x_270); -if (lean_obj_tag(x_275) == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = lean_st_ref_take(x_5, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_array_push(x_15, x_1); +x_17 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_2); +x_18 = lean_st_ref_set(x_5, x_17, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_276; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_273); -x_276 = l_Lean_Meta_isProp(x_273, x_9, x_10, x_11, x_12, x_274); -if (lean_obj_tag(x_276) == 0) +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else { -lean_object* x_277; uint8_t x_278; -x_277 = lean_ctor_get(x_276, 0); -lean_inc(x_277); -x_278 = lean_unbox(x_277); -lean_dec(x_277); -if (x_278 == 0) +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1() { +_start: { -lean_object* x_279; lean_object* x_280; -x_279 = lean_ctor_get(x_276, 1); -lean_inc(x_279); -lean_dec(x_276); -lean_inc(x_9); -lean_inc(x_270); -x_280 = l_Lean_FVarId_getDecl(x_270, x_9, x_10, x_11, x_12, x_279); -if (lean_obj_tag(x_280) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found full coverage", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2() { +_start: { -lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; -x_281 = lean_ctor_get(x_280, 0); -lean_inc(x_281); -x_282 = lean_ctor_get(x_280, 1); -lean_inc(x_282); -lean_dec(x_280); -x_283 = l_Lean_LocalDecl_binderInfo(x_281); -lean_dec(x_281); -x_284 = lean_box(x_283); -if (lean_obj_tag(x_284) == 3) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_285; -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = lean_st_ref_get(x_6, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 2); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +x_18 = lean_array_get_size(x_17); +lean_dec(x_17); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_273); -lean_inc(x_266); -lean_inc(x_3); -x_285 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized(x_3, x_266, x_273, x_9, x_10, x_11, x_12, x_282); -if (lean_obj_tag(x_285) == 0) +lean_inc(x_8); +lean_inc(x_7); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(x_16, x_18, x_15, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_286; uint8_t x_287; -x_286 = lean_ctor_get(x_285, 0); -lean_inc(x_286); -x_287 = lean_unbox(x_286); -lean_dec(x_286); -if (x_287 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; -lean_dec(x_273); -lean_dec(x_270); -x_288 = lean_ctor_get(x_285, 1); -lean_inc(x_288); -lean_dec(x_285); -if (lean_is_scalar(x_269)) { - x_289 = lean_alloc_ctor(0, 2, 0); -} else { - x_289 = x_269; -} -lean_ctor_set(x_289, 0, x_267); -lean_ctor_set(x_289, 1, x_268); -x_290 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_290, 0, x_266); -lean_ctor_set(x_290, 1, x_289); -x_291 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_291, 0, x_290); -x_17 = x_291; -x_18 = x_288; -goto block_23; +uint8_t x_38; +x_38 = 1; +x_22 = x_38; +goto block_37; } else { -lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; uint8_t x_303; lean_object* x_304; -lean_dec(x_267); -x_292 = lean_ctor_get(x_285, 1); -lean_inc(x_292); -lean_dec(x_285); -x_293 = lean_box(0); -lean_inc(x_270); -x_294 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_293); -x_295 = lean_box(0); -lean_inc(x_2); -x_296 = lean_array_mk(x_2); -x_297 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_298 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_298, 0, x_297); -lean_ctor_set(x_298, 1, x_295); -lean_ctor_set(x_298, 2, x_296); -x_299 = l_Lean_CollectFVars_main(x_273, x_298); -x_300 = lean_ctor_get(x_299, 2); -lean_inc(x_300); -lean_dec(x_299); -x_301 = lean_array_get_size(x_300); -x_302 = lean_unsigned_to_nat(0u); -x_303 = lean_nat_dec_lt(x_302, x_301); -x_304 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_293); -if (x_303 == 0) -{ -uint8_t x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; -lean_dec(x_301); -lean_dec(x_300); -x_305 = 1; -x_306 = lean_box(x_305); -if (lean_is_scalar(x_269)) { - x_307 = lean_alloc_ctor(0, 2, 0); -} else { - x_307 = x_269; -} -lean_ctor_set(x_307, 0, x_306); -lean_ctor_set(x_307, 1, x_304); -x_308 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_308, 0, x_294); -lean_ctor_set(x_308, 1, x_307); -x_309 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_309, 0, x_308); -x_17 = x_309; -x_18 = x_292; -goto block_23; +uint8_t x_39; +lean_dec(x_20); +x_39 = 0; +x_22 = x_39; +goto block_37; } -else -{ -uint8_t x_310; -x_310 = lean_nat_dec_le(x_301, x_301); -if (x_310 == 0) +block_37: { -uint8_t x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; -lean_dec(x_301); -lean_dec(x_300); -x_311 = 1; -x_312 = lean_box(x_311); -if (lean_is_scalar(x_269)) { - x_313 = lean_alloc_ctor(0, 2, 0); -} else { - x_313 = x_269; -} -lean_ctor_set(x_313, 0, x_312); -lean_ctor_set(x_313, 1, x_304); -x_314 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_314, 0, x_294); -lean_ctor_set(x_314, 1, x_313); -x_315 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_315, 0, x_314); -x_17 = x_315; -x_18 = x_292; -goto block_23; +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_2); +x_23 = lean_box(0); +x_24 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(x_1, x_22, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +return x_24; } else { -size_t x_316; size_t x_317; lean_object* x_318; uint8_t x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; -x_316 = 0; -x_317 = lean_usize_of_nat(x_301); -lean_dec(x_301); -x_318 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_3, x_300, x_316, x_317, x_294); -lean_dec(x_300); -x_319 = 1; -x_320 = lean_box(x_319); -if (lean_is_scalar(x_269)) { - x_321 = lean_alloc_ctor(0, 2, 0); -} else { - x_321 = x_269; +lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_inc(x_2); +x_25 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_unbox(x_26); +lean_dec(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_2); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_box(0); +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(x_1, x_22, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +return x_30; } -lean_ctor_set(x_321, 0, x_320); -lean_ctor_set(x_321, 1, x_304); -x_322 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_322, 0, x_318); -lean_ctor_set(x_322, 1, x_321); -x_323 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_323, 0, x_322); -x_17 = x_323; -x_18 = x_292; -goto block_23; +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2; +x_33 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_2, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(x_1, x_22, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_34); +return x_36; } } } } else { -lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; -lean_dec(x_273); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_266); -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_40; lean_dec(x_10); lean_dec(x_9); -lean_dec(x_3); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); lean_dec(x_2); -x_324 = lean_ctor_get(x_285, 0); -lean_inc(x_324); -x_325 = lean_ctor_get(x_285, 1); -lean_inc(x_325); -if (lean_is_exclusive(x_285)) { - lean_ctor_release(x_285, 0); - lean_ctor_release(x_285, 1); - x_326 = x_285; -} else { - lean_dec_ref(x_285); - x_326 = lean_box(0); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_19); +if (x_40 == 0) +{ +return x_19; } -if (lean_is_scalar(x_326)) { - x_327 = lean_alloc_ctor(1, 2, 0); -} else { - x_327 = x_326; +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_19, 0); +x_42 = lean_ctor_get(x_19, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_19); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } -lean_ctor_set(x_327, 0, x_324); -lean_ctor_set(x_327, 1, x_325); -return x_327; } } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1() { +_start: { -lean_object* x_328; lean_object* x_329; lean_object* x_330; -lean_dec(x_284); -lean_dec(x_273); -lean_dec(x_270); -if (lean_is_scalar(x_269)) { - x_328 = lean_alloc_ctor(0, 2, 0); -} else { - x_328 = x_269; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("search", 6, 6); +return x_1; } -lean_ctor_set(x_328, 0, x_267); -lean_ctor_set(x_328, 1, x_268); -x_329 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_329, 0, x_266); -lean_ctor_set(x_329, 1, x_328); -x_330 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_330, 0, x_329); -x_17 = x_330; -x_18 = x_282; -goto block_23; } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_2 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1; +x_3 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3() { +_start: { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; -lean_dec(x_273); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_266); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_331 = lean_ctor_get(x_280, 0); -lean_inc(x_331); -x_332 = lean_ctor_get(x_280, 1); -lean_inc(x_332); -if (lean_is_exclusive(x_280)) { - lean_ctor_release(x_280, 0); - lean_ctor_release(x_280, 1); - x_333 = x_280; -} else { - lean_dec_ref(x_280); - x_333 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found pattern: ", 15, 15); +return x_1; } -if (lean_is_scalar(x_333)) { - x_334 = lean_alloc_ctor(1, 2, 0); -} else { - x_334 = x_333; } -lean_ctor_set(x_334, 0, x_331); -lean_ctor_set(x_334, 1, x_332); -return x_334; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_335; uint8_t x_336; -x_335 = lean_ctor_get(x_276, 1); -lean_inc(x_335); -lean_dec(x_276); -x_336 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkTypeFVars(x_3, x_266, x_273); -if (x_336 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2; +x_11 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) { -lean_object* x_337; lean_object* x_338; lean_object* x_339; -lean_dec(x_270); -if (lean_is_scalar(x_269)) { - x_337 = lean_alloc_ctor(0, 2, 0); -} else { - x_337 = x_269; -} -lean_ctor_set(x_337, 0, x_267); -lean_ctor_set(x_337, 1, x_268); -x_338 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_338, 0, x_266); -lean_ctor_set(x_338, 1, x_337); -x_339 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_339, 0, x_338); -x_17 = x_339; -x_18 = x_335; -goto block_23; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = lean_box(0); +x_16 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(x_1, x_10, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_14); +return x_16; } else { -lean_object* x_340; lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; -lean_dec(x_267); -x_340 = lean_box(0); -lean_inc(x_270); -x_341 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_266, x_270, x_340); -x_342 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_340); -x_343 = 1; -x_344 = lean_box(x_343); -if (lean_is_scalar(x_269)) { - x_345 = lean_alloc_ctor(0, 2, 0); -} else { - x_345 = x_269; +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_18 = lean_ctor_get(x_11, 1); +x_19 = lean_ctor_get(x_11, 0); +lean_dec(x_19); +lean_inc(x_1); +x_20 = l_Lean_Meta_Grind_ppPattern(x_1); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4; +lean_ctor_set_tag(x_11, 7); +lean_ctor_set(x_11, 1, x_20); +lean_ctor_set(x_11, 0, x_21); +x_22 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_11); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_10, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(x_1, x_10, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_26); +lean_dec(x_25); +return x_27; } -lean_ctor_set(x_345, 0, x_344); -lean_ctor_set(x_345, 1, x_342); -x_346 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_346, 0, x_341); -lean_ctor_set(x_346, 1, x_345); -x_347 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_347, 0, x_346); -x_17 = x_347; -x_18 = x_335; -goto block_23; +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +lean_dec(x_11); +lean_inc(x_1); +x_29 = l_Lean_Meta_Grind_ppPattern(x_1); +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_10, x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(x_1, x_10, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_36); +lean_dec(x_35); +return x_37; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; -lean_dec(x_273); -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_266); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_10; +x_10 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_348 = lean_ctor_get(x_276, 0); -lean_inc(x_348); -x_349 = lean_ctor_get(x_276, 1); -lean_inc(x_349); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - lean_ctor_release(x_276, 1); - x_350 = x_276; -} else { - lean_dec_ref(x_276); - x_350 = lean_box(0); +return x_10; } -if (lean_is_scalar(x_350)) { - x_351 = lean_alloc_ctor(1, 2, 0); -} else { - x_351 = x_350; } -lean_ctor_set(x_351, 0, x_348); -lean_ctor_set(x_351, 1, x_349); -return x_351; +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; uint8_t x_361; lean_object* x_362; lean_object* x_363; -lean_dec(x_267); -if (lean_is_exclusive(x_275)) { - lean_ctor_release(x_275, 0); - x_352 = x_275; -} else { - lean_dec_ref(x_275); - x_352 = lean_box(0); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_2); +lean_dec(x_2); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__1(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_13; } -x_353 = lean_box(0); -lean_inc(x_2); -x_354 = lean_array_mk(x_2); -x_355 = l_Lean_Meta_Grind_NormalizePattern_main___closed__4; -x_356 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_356, 0, x_355); -lean_ctor_set(x_356, 1, x_353); -lean_ctor_set(x_356, 2, x_354); -x_357 = l_Lean_CollectFVars_main(x_273, x_356); -x_358 = lean_ctor_get(x_357, 2); -lean_inc(x_358); -lean_dec(x_357); -x_359 = lean_array_get_size(x_358); -x_360 = lean_unsigned_to_nat(0u); -x_361 = lean_nat_dec_lt(x_360, x_359); -x_362 = lean_box(0); -x_363 = l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(x_268, x_270, x_362); -if (x_361 == 0) +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; -lean_dec(x_359); -lean_dec(x_358); -x_364 = 1; -x_365 = lean_box(x_364); -if (lean_is_scalar(x_269)) { - x_366 = lean_alloc_ctor(0, 2, 0); -} else { - x_366 = x_269; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_12; } -lean_ctor_set(x_366, 0, x_365); -lean_ctor_set(x_366, 1, x_363); -x_367 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_367, 0, x_266); -lean_ctor_set(x_367, 1, x_366); -if (lean_is_scalar(x_352)) { - x_368 = lean_alloc_ctor(1, 1, 0); -} else { - x_368 = x_352; } -lean_ctor_set(x_368, 0, x_367); -x_17 = x_368; -x_18 = x_274; -goto block_23; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +return x_10; } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -uint8_t x_369; -x_369 = lean_nat_dec_le(x_359, x_359); -if (x_369 == 0) +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_5, x_4); +if (x_15 == 0) { -uint8_t x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; -lean_dec(x_359); -lean_dec(x_358); -x_370 = 1; -x_371 = lean_box(x_370); -if (lean_is_scalar(x_269)) { - x_372 = lean_alloc_ctor(0, 2, 0); -} else { - x_372 = x_269; -} -lean_ctor_set(x_372, 0, x_371); -lean_ctor_set(x_372, 1, x_363); -x_373 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_373, 0, x_266); -lean_ctor_set(x_373, 1, x_372); -if (lean_is_scalar(x_352)) { - x_374 = lean_alloc_ctor(1, 1, 0); -} else { - x_374 = x_352; +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_6); +lean_ctor_set(x_16, 1, x_14); +return x_16; } -lean_ctor_set(x_374, 0, x_373); -x_17 = x_374; -x_18 = x_274; -goto block_23; +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_array_uget(x_3, x_5); +x_18 = lean_ctor_get(x_6, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_6, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_6, 2); +lean_inc(x_20); +x_21 = lean_nat_dec_lt(x_19, x_20); +if (x_21 == 0) +{ +lean_object* x_22; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_6); +lean_ctor_set(x_22, 1, x_14); +return x_22; } else { -size_t x_375; size_t x_376; lean_object* x_377; uint8_t x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; -x_375 = 0; -x_376 = lean_usize_of_nat(x_359); -lean_dec(x_359); -x_377 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_3, x_358, x_375, x_376, x_266); -lean_dec(x_358); -x_378 = 1; -x_379 = lean_box(x_378); -if (lean_is_scalar(x_269)) { - x_380 = lean_alloc_ctor(0, 2, 0); -} else { - x_380 = x_269; +uint8_t x_23; +x_23 = !lean_is_exclusive(x_6); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_6, 2); +lean_dec(x_24); +x_25 = lean_ctor_get(x_6, 1); +lean_dec(x_25); +x_26 = lean_ctor_get(x_6, 0); +lean_dec(x_26); +x_27 = lean_array_fget(x_18, x_19); +x_28 = lean_unbox(x_27); +lean_dec(x_27); +x_29 = lean_unsigned_to_nat(1u); +x_30 = lean_nat_add(x_19, x_29); +lean_dec(x_19); +lean_ctor_set(x_6, 1, x_30); +if (x_28 == 0) +{ +lean_object* x_31; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; size_t x_33; size_t x_34; +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = 1; +x_34 = lean_usize_add(x_5, x_33); +x_5 = x_34; +x_14 = x_32; +goto _start; } -lean_ctor_set(x_380, 0, x_379); -lean_ctor_set(x_380, 1, x_363); -x_381 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_381, 0, x_377); -lean_ctor_set(x_381, 1, x_380); -if (lean_is_scalar(x_352)) { - x_382 = lean_alloc_ctor(1, 1, 0); -} else { - x_382 = x_352; +else +{ +uint8_t x_36; +lean_dec(x_6); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) +{ +return x_31; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } -lean_ctor_set(x_382, 0, x_381); -x_17 = x_382; -x_18 = x_274; -goto block_23; } } +else +{ +size_t x_40; size_t x_41; +lean_dec(x_17); +x_40 = 1; +x_41 = lean_usize_add(x_5, x_40); +x_5 = x_41; +goto _start; } } else { -lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; -lean_dec(x_270); -lean_dec(x_269); -lean_dec(x_268); -lean_dec(x_267); -lean_dec(x_266); +lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_6); +x_43 = lean_array_fget(x_18, x_19); +x_44 = lean_unbox(x_43); +lean_dec(x_43); +x_45 = lean_unsigned_to_nat(1u); +x_46 = lean_nat_add(x_19, x_45); +lean_dec(x_19); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_18); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_20); +if (x_44 == 0) +{ +lean_object* x_48; +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_48 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; size_t x_50; size_t x_51; +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = 1; +x_51 = lean_usize_add(x_5, x_50); +x_5 = x_51; +x_6 = x_47; +x_14 = x_49; +goto _start; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_47); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_2); -x_383 = lean_ctor_get(x_272, 0); -lean_inc(x_383); -x_384 = lean_ctor_get(x_272, 1); -lean_inc(x_384); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - lean_ctor_release(x_272, 1); - x_385 = x_272; +lean_dec(x_8); +lean_dec(x_7); +x_53 = lean_ctor_get(x_48, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_48, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_55 = x_48; } else { - lean_dec_ref(x_272); - x_385 = lean_box(0); + lean_dec_ref(x_48); + x_55 = lean_box(0); } -if (lean_is_scalar(x_385)) { - x_386 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_55)) { + x_56 = lean_alloc_ctor(1, 2, 0); } else { - x_386 = x_385; + x_56 = x_55; } -lean_ctor_set(x_386, 0, x_383); -lean_ctor_set(x_386, 1, x_384); -return x_386; +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_54); +return x_56; } } else { -lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; -lean_dec(x_270); -lean_dec(x_16); -if (lean_is_exclusive(x_271)) { - lean_ctor_release(x_271, 0); - x_387 = x_271; -} else { - lean_dec_ref(x_271); - x_387 = lean_box(0); -} -if (lean_is_scalar(x_269)) { - x_388 = lean_alloc_ctor(0, 2, 0); -} else { - x_388 = x_269; -} -lean_ctor_set(x_388, 0, x_267); -lean_ctor_set(x_388, 1, x_268); -x_389 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_389, 0, x_266); -lean_ctor_set(x_389, 1, x_388); -if (lean_is_scalar(x_387)) { - x_390 = lean_alloc_ctor(1, 1, 0); -} else { - x_390 = x_387; -} -lean_ctor_set(x_390, 0, x_389); -x_17 = x_390; -x_18 = x_13; -goto block_23; -} -} -block_23: -{ -lean_object* x_19; size_t x_20; size_t x_21; -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); +size_t x_57; size_t x_58; lean_dec(x_17); -x_20 = 1; -x_21 = lean_usize_add(x_7, x_20); -x_7 = x_21; -x_8 = x_19; -x_13 = x_18; +x_57 = 1; +x_58 = lean_usize_add(x_5, x_57); +x_5 = x_58; +x_6 = x_47; goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (x_1 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_11 = lean_box(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_3); -lean_ctor_set(x_13, 1, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_4); -lean_ctor_set(x_14, 1, x_13); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_10); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_box(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_2); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_3); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_4); -lean_ctor_set(x_20, 1, x_19); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_10); -return x_22; -} -} -} -static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_7, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_12); +x_14 = l_Lean_Meta_Grind_ppPattern___closed__5; lean_inc(x_13); -lean_dec(x_7); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_13, 1); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; -x_17 = lean_ctor_get(x_15, 0); -lean_dec(x_17); -x_18 = lean_box(0); -x_19 = 0; -x_20 = lean_box(x_19); -lean_ctor_set(x_15, 0, x_20); -x_21 = lean_array_size(x_2); -x_22 = 0; -lean_inc(x_11); +x_15 = lean_mk_array(x_13, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_13, x_16); +lean_dec(x_13); +x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_15, x_17); +x_19 = lean_array_get_size(x_18); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_5); -lean_inc(x_3); -x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_18, x_2, x_21, x_22, x_13, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_23) == 0) +lean_inc(x_7); +x_20 = l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask(x_2, x_19, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -x_26 = !lean_is_exclusive(x_23); -if (x_26 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_array_get_size(x_21); +x_24 = l_Array_toSubarray___rarg(x_21, x_12, x_23); +x_25 = lean_box(0); +x_26 = lean_array_size(x_18); +x_27 = 0; +x_28 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1(x_18, x_25, x_18, x_26, x_27, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +lean_dec(x_18); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_23, 1); -x_28 = lean_ctor_get(x_23, 0); -lean_dec(x_28); -x_29 = !lean_is_exclusive(x_24); +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); if (x_29 == 0) { -lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = lean_ctor_get(x_24, 0); -x_31 = lean_ctor_get(x_24, 1); -lean_dec(x_31); -x_32 = !lean_is_exclusive(x_25); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_25, 0); -x_34 = lean_ctor_get(x_25, 1); -x_35 = lean_unsigned_to_nat(0u); -x_36 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_35, x_30); -x_37 = lean_nat_dec_eq(x_36, x_1); -lean_dec(x_36); -if (x_37 == 0) -{ -lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; -lean_free_object(x_25); -lean_free_object(x_24); -lean_free_object(x_23); -x_38 = lean_box(0); -x_39 = lean_unbox(x_33); -lean_dec(x_33); -lean_inc(x_6); -x_40 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_39, x_34, x_30, x_6, x_38, x_8, x_9, x_10, x_11, x_27); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) -{ -uint8_t x_42; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_42 = !lean_is_exclusive(x_40); -if (x_42 == 0) +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 0); +lean_dec(x_30); +x_31 = lean_box(0); +lean_ctor_set(x_28, 0, x_31); +return x_28; +} +else { -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_40, 0); -lean_dec(x_43); -x_44 = lean_ctor_get(x_41, 0); -lean_inc(x_44); -lean_dec(x_41); -lean_ctor_set(x_40, 0, x_44); -return x_40; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_28, 1); +lean_inc(x_32); +lean_dec(x_28); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_40, 1); -lean_inc(x_45); -lean_dec(x_40); -x_46 = lean_ctor_get(x_41, 0); -lean_inc(x_46); -lean_dec(x_41); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_45); -return x_47; -} +uint8_t x_35; +x_35 = !lean_is_exclusive(x_28); +if (x_35 == 0) +{ +return x_28; } else { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_40, 1); -lean_inc(x_48); -lean_dec(x_40); -x_49 = lean_ctor_get(x_41, 0); -lean_inc(x_49); -lean_dec(x_41); -x_7 = x_49; -x_12 = x_48; -goto _start; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_28, 0); +x_37 = lean_ctor_get(x_28, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_28); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} } } else { -lean_object* x_51; lean_object* x_52; -lean_dec(x_11); +uint8_t x_39; +lean_dec(x_18); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -x_51 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_24); -lean_ctor_set(x_23, 0, x_52); -return x_23; -} +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_20); +if (x_39 == 0) +{ +return x_20; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_53 = lean_ctor_get(x_25, 0); -x_54 = lean_ctor_get(x_25, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_25); -x_55 = lean_unsigned_to_nat(0u); -x_56 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_55, x_30); -x_57 = lean_nat_dec_eq(x_56, x_1); -lean_dec(x_56); -if (x_57 == 0) -{ -lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; -lean_free_object(x_24); -lean_free_object(x_23); -x_58 = lean_box(0); -x_59 = lean_unbox(x_53); -lean_dec(x_53); -lean_inc(x_6); -x_60 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_59, x_54, x_30, x_6, x_58, x_8, x_9, x_10, x_11, x_27); -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -if (lean_obj_tag(x_61) == 0) -{ -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -if (lean_is_exclusive(x_60)) { - lean_ctor_release(x_60, 0); - lean_ctor_release(x_60, 1); - x_63 = x_60; -} else { - lean_dec_ref(x_60); - x_63 = lean_box(0); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_20, 0); +x_41 = lean_ctor_get(x_20, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_20); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } -x_64 = lean_ctor_get(x_61, 0); -lean_inc(x_64); -lean_dec(x_61); -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_63; } -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_62); -return x_65; } -else -{ -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_60, 1); -lean_inc(x_66); -lean_dec(x_60); -x_67 = lean_ctor_get(x_61, 0); -lean_inc(x_67); -lean_dec(x_61); -x_7 = x_67; -x_12 = x_66; -goto _start; } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_set(x_5, x_1, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_11, 0, x_15); +return x_11; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_11, 0); +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_inc(x_16); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_53); -lean_ctor_set(x_69, 1, x_54); -lean_ctor_set(x_24, 1, x_69); -x_70 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_24); -lean_ctor_set(x_23, 0, x_71); -return x_23; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; } } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1() { +_start: { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_72 = lean_ctor_get(x_24, 0); -lean_inc(x_72); -lean_dec(x_24); -x_73 = lean_ctor_get(x_25, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_25, 1); -lean_inc(x_74); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_75 = x_25; -} else { - lean_dec_ref(x_25); - x_75 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("skip, no new variables covered", 30, 30); +return x_1; } -x_76 = lean_unsigned_to_nat(0u); -x_77 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_76, x_72); -x_78 = lean_nat_dec_eq(x_77, x_1); -lean_dec(x_77); -if (x_78 == 0) +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2() { +_start: { -lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; -lean_dec(x_75); -lean_free_object(x_23); -x_79 = lean_box(0); -x_80 = lean_unbox(x_73); -lean_dec(x_73); -lean_inc(x_6); -x_81 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_80, x_74, x_72, x_6, x_79, x_8, x_9, x_10, x_11, x_27); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -if (lean_obj_tag(x_82) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_84 = x_81; -} else { - lean_dec_ref(x_81); - x_84 = lean_box(0); +lean_object* x_12; lean_object* x_13; uint8_t x_14; +lean_inc(x_1); +x_12 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_1); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2; +x_20 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_2, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +lean_dec(x_21); +return x_23; } -x_85 = lean_ctor_get(x_82, 0); -lean_inc(x_85); -lean_dec(x_82); -if (lean_is_scalar(x_84)) { - x_86 = lean_alloc_ctor(0, 2, 0); -} else { - x_86 = x_84; } -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_83); -return x_86; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1() { +_start: { -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_81, 1); -lean_inc(x_87); -lean_dec(x_81); -x_88 = lean_ctor_get(x_82, 0); -lean_inc(x_88); -lean_dec(x_82); -x_7 = x_88; -x_12 = x_87; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +uint8_t x_13; lean_object* x_14; +x_13 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go(x_1, x_13, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_st_ref_get(x_7, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_3, 2); +lean_inc(x_20); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_ctor_get(x_18, 2); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_nat_dec_lt(x_21, x_23); +lean_dec(x_23); +lean_dec(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_15); +x_25 = lean_box(0); +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3(x_2, x_3, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); lean_dec(x_5); +return x_26; +} +else +{ +lean_object* x_27; lean_dec(x_3); -if (lean_is_scalar(x_75)) { - x_90 = lean_alloc_ctor(0, 2, 0); -} else { - x_90 = x_75; +lean_dec(x_2); +x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +lean_dec(x_7); +if (lean_obj_tag(x_27) == 0) +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1; +lean_ctor_set(x_27, 0, x_30); +return x_27; } -lean_ctor_set(x_90, 0, x_73); -lean_ctor_set(x_90, 1, x_74); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_72); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_91); -lean_ctor_set(x_23, 0, x_93); -return x_23; +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_27); +if (x_34 == 0) +{ +return x_27; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; -x_94 = lean_ctor_get(x_23, 1); -lean_inc(x_94); -lean_dec(x_23); -x_95 = lean_ctor_get(x_24, 0); -lean_inc(x_95); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_96 = x_24; -} else { - lean_dec_ref(x_24); - x_96 = lean_box(0); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_27, 0); +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_27); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } -x_97 = lean_ctor_get(x_25, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_25, 1); -lean_inc(x_98); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_99 = x_25; -} else { - lean_dec_ref(x_25); - x_99 = lean_box(0); } -x_100 = lean_unsigned_to_nat(0u); -x_101 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_100, x_95); -x_102 = lean_nat_dec_eq(x_101, x_1); -lean_dec(x_101); -if (x_102 == 0) -{ -lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_99); -lean_dec(x_96); -x_103 = lean_box(0); -x_104 = lean_unbox(x_97); -lean_dec(x_97); -lean_inc(x_6); -x_105 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_104, x_98, x_95, x_6, x_103, x_8, x_9, x_10, x_11, x_94); -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -if (lean_obj_tag(x_106) == 0) +} +} +else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +uint8_t x_38; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_3); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - x_108 = x_105; -} else { - lean_dec_ref(x_105); - x_108 = lean_box(0); +lean_dec(x_2); +x_38 = !lean_is_exclusive(x_14); +if (x_38 == 0) +{ +return x_14; } -x_109 = lean_ctor_get(x_106, 0); -lean_inc(x_109); -lean_dec(x_106); -if (lean_is_scalar(x_108)) { - x_110 = lean_alloc_ctor(0, 2, 0); -} else { - x_110 = x_108; +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_14, 0); +x_40 = lean_ctor_get(x_14, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_14); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_107); -return x_110; } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_111; lean_object* x_112; -x_111 = lean_ctor_get(x_105, 1); -lean_inc(x_111); -lean_dec(x_105); -x_112 = lean_ctor_get(x_106, 0); -lean_inc(x_112); -lean_dec(x_106); -x_7 = x_112; -x_12 = x_111; -goto _start; +lean_object* x_10; lean_object* x_11; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1() { +_start: { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -if (lean_is_scalar(x_99)) { - x_114 = lean_alloc_ctor(0, 2, 0); -} else { - x_114 = x_99; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5___boxed), 9, 0); +return x_1; } -lean_ctor_set(x_114, 0, x_97); -lean_ctor_set(x_114, 1, x_98); -if (lean_is_scalar(x_96)) { - x_115 = lean_alloc_ctor(0, 2, 0); -} else { - x_115 = x_96; } -lean_ctor_set(x_115, 0, x_95); -lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_117 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_115); -x_118 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_94); -return x_118; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("skip, does not contain pattern variables", 40, 40); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_119; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +x_14 = lean_expr_abstract(x_1, x_13); +lean_dec(x_13); +x_15 = l_Lean_Expr_hasLooseBVars(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +lean_dec(x_14); lean_dec(x_3); -x_119 = !lean_is_exclusive(x_23); -if (x_119 == 0) +lean_inc(x_2); +x_16 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1; +x_20 = lean_unbox(x_17); +lean_dec(x_17); +if (x_20 == 0) { -return x_23; +lean_object* x_21; lean_object* x_22; +lean_dec(x_2); +x_21 = lean_box(0); +x_22 = lean_apply_9(x_19, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_22; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_120 = lean_ctor_get(x_23, 0); -x_121 = lean_ctor_get(x_23, 1); -lean_inc(x_121); -lean_inc(x_120); -lean_dec(x_23); -x_122 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -return x_122; -} +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3; +x_24 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_2, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_apply_9(x_19, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +return x_27; } } else { -lean_object* x_123; lean_object* x_124; uint8_t x_125; lean_object* x_126; lean_object* x_127; size_t x_128; size_t x_129; lean_object* x_130; -x_123 = lean_ctor_get(x_15, 1); -lean_inc(x_123); -lean_dec(x_15); -x_124 = lean_box(0); -x_125 = 0; -x_126 = lean_box(x_125); -x_127 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_127, 0, x_126); -lean_ctor_set(x_127, 1, x_123); -lean_ctor_set(x_13, 1, x_127); -x_128 = lean_array_size(x_2); -x_129 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_5); -lean_inc(x_3); -x_130 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_124, x_2, x_128, x_129, x_13, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_130) == 0) -{ -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_131, 1); -lean_inc(x_132); -x_133 = lean_ctor_get(x_130, 1); -lean_inc(x_133); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_134 = x_130; -} else { - lean_dec_ref(x_130); - x_134 = lean_box(0); +lean_object* x_28; lean_object* x_29; +x_28 = lean_box(0); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4(x_14, x_2, x_3, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_6); +return x_29; } -x_135 = lean_ctor_get(x_131, 0); -lean_inc(x_135); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_136 = x_131; -} else { - lean_dec_ref(x_131); - x_136 = lean_box(0); } -x_137 = lean_ctor_get(x_132, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_132, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_132)) { - lean_ctor_release(x_132, 0); - lean_ctor_release(x_132, 1); - x_139 = x_132; -} else { - lean_dec_ref(x_132); - x_139 = lean_box(0); } -x_140 = lean_unsigned_to_nat(0u); -x_141 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_140, x_135); -x_142 = lean_nat_dec_eq(x_141, x_1); -lean_dec(x_141); -if (x_142 == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1() { +_start: { -lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_139); -lean_dec(x_136); -lean_dec(x_134); -x_143 = lean_box(0); -x_144 = lean_unbox(x_137); -lean_dec(x_137); -lean_inc(x_6); -x_145 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_144, x_138, x_135, x_6, x_143, x_8, x_9, x_10, x_11, x_133); -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -if (lean_obj_tag(x_146) == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("skip, exception during normalization", 36, 36); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2() { +_start: { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_145)) { - lean_ctor_release(x_145, 0); - lean_ctor_release(x_145, 1); - x_148 = x_145; -} else { - lean_dec_ref(x_145); - x_148 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -x_149 = lean_ctor_get(x_146, 0); -lean_inc(x_149); -lean_dec(x_146); -if (lean_is_scalar(x_148)) { - x_150 = lean_alloc_ctor(0, 2, 0); -} else { - x_150 = x_148; } -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_147); -return x_150; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("candidate: ", 11, 11); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4() { +_start: { -lean_object* x_151; lean_object* x_152; -x_151 = lean_ctor_get(x_145, 1); -lean_inc(x_151); -lean_dec(x_145); -x_152 = lean_ctor_get(x_146, 0); -lean_inc(x_152); -lean_dec(x_146); -x_7 = x_152; -x_12 = x_151; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -if (lean_is_scalar(x_139)) { - x_154 = lean_alloc_ctor(0, 2, 0); -} else { - x_154 = x_139; +switch (lean_obj_tag(x_1)) { +case 5: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l_Lean_Expr_getAppFn(x_1); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_isPatternFnCandidate(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(x_1, x_11, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_17; } -lean_ctor_set(x_154, 0, x_137); -lean_ctor_set(x_154, 1, x_138); -if (lean_is_scalar(x_136)) { - x_155 = lean_alloc_ctor(0, 2, 0); +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_32; lean_object* x_33; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + x_19 = x_12; } else { - x_155 = x_136; + lean_dec_ref(x_12); + x_19 = lean_box(0); } -lean_ctor_set(x_155, 0, x_135); -lean_ctor_set(x_155, 1, x_154); -x_156 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_157 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_155); -if (lean_is_scalar(x_134)) { - x_158 = lean_alloc_ctor(0, 2, 0); +x_20 = lean_st_ref_get(x_5, x_18); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + lean_ctor_release(x_20, 1); + x_23 = x_20; } else { - x_158 = x_134; -} -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_133); -return x_158; + lean_dec_ref(x_20); + x_23 = lean_box(0); } +x_24 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2; +x_55 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_unbox(x_56); +lean_dec(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_55, 1); +lean_inc(x_58); +lean_dec(x_55); +x_59 = lean_box(0); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_21); +x_60 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(x_1, x_24, x_21, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_58); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; +lean_dec(x_21); +lean_dec(x_19); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_25 = x_61; +x_26 = x_62; +goto block_31; } else { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_159 = lean_ctor_get(x_130, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_130, 1); -lean_inc(x_160); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_161 = x_130; -} else { - lean_dec_ref(x_130); - x_161 = lean_box(0); -} -if (lean_is_scalar(x_161)) { - x_162 = lean_alloc_ctor(1, 2, 0); -} else { - x_162 = x_161; -} -lean_ctor_set(x_162, 0, x_159); -lean_ctor_set(x_162, 1, x_160); -return x_162; +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_60, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_60, 1); +lean_inc(x_64); +lean_dec(x_60); +x_32 = x_63; +x_33 = x_64; +goto block_54; } } +else +{ +uint8_t x_65; +x_65 = !lean_is_exclusive(x_55); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_66 = lean_ctor_get(x_55, 1); +x_67 = lean_ctor_get(x_55, 0); +lean_dec(x_67); +lean_inc(x_1); +x_68 = l_Lean_MessageData_ofExpr(x_1); +x_69 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4; +lean_ctor_set_tag(x_55, 7); +lean_ctor_set(x_55, 1, x_68); +lean_ctor_set(x_55, 0, x_69); +x_70 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_71 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_71, 0, x_55); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_24, x_71, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_66); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_21); +x_75 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(x_1, x_24, x_21, x_73, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_74); +lean_dec(x_73); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; +lean_dec(x_21); +lean_dec(x_19); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec(x_75); +x_25 = x_76; +x_26 = x_77; +goto block_31; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; size_t x_172; size_t x_173; lean_object* x_174; -x_163 = lean_ctor_get(x_13, 1); -x_164 = lean_ctor_get(x_13, 0); -lean_inc(x_163); -lean_inc(x_164); -lean_dec(x_13); -x_165 = lean_ctor_get(x_163, 1); -lean_inc(x_165); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_166 = x_163; -} else { - lean_dec_ref(x_163); - x_166 = lean_box(0); +lean_object* x_78; lean_object* x_79; +x_78 = lean_ctor_get(x_75, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_75, 1); +lean_inc(x_79); +lean_dec(x_75); +x_32 = x_78; +x_33 = x_79; +goto block_54; } -x_167 = lean_box(0); -x_168 = 0; -x_169 = lean_box(x_168); -if (lean_is_scalar(x_166)) { - x_170 = lean_alloc_ctor(0, 2, 0); -} else { - x_170 = x_166; } -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_165); -x_171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_171, 0, x_164); -lean_ctor_set(x_171, 1, x_170); -x_172 = lean_array_size(x_2); -x_173 = 0; -lean_inc(x_11); -lean_inc(x_10); +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_80 = lean_ctor_get(x_55, 1); +lean_inc(x_80); +lean_dec(x_55); +lean_inc(x_1); +x_81 = l_Lean_MessageData_ofExpr(x_1); +x_82 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4; +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_81); +x_84 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_24, x_85, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_80); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); lean_inc(x_9); lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); +lean_inc(x_4); lean_inc(x_3); -x_174 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_5, x_167, x_2, x_172, x_173, x_171, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_174) == 0) +lean_inc(x_21); +x_89 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(x_1, x_24, x_21, x_87, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_88); +lean_dec(x_87); +if (lean_obj_tag(x_89) == 0) { -lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; -x_175 = lean_ctor_get(x_174, 0); -lean_inc(x_175); -x_176 = lean_ctor_get(x_175, 1); -lean_inc(x_176); -x_177 = lean_ctor_get(x_174, 1); -lean_inc(x_177); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_178 = x_174; -} else { - lean_dec_ref(x_174); - x_178 = lean_box(0); +lean_object* x_90; lean_object* x_91; +lean_dec(x_21); +lean_dec(x_19); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); +x_25 = x_90; +x_26 = x_91; +goto block_31; } -x_179 = lean_ctor_get(x_175, 0); -lean_inc(x_179); -if (lean_is_exclusive(x_175)) { - lean_ctor_release(x_175, 0); - lean_ctor_release(x_175, 1); - x_180 = x_175; -} else { - lean_dec_ref(x_175); - x_180 = lean_box(0); +else +{ +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_89, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_89, 1); +lean_inc(x_93); +lean_dec(x_89); +x_32 = x_92; +x_33 = x_93; +goto block_54; } -x_181 = lean_ctor_get(x_176, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_176, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_183 = x_176; -} else { - lean_dec_ref(x_176); - x_183 = lean_box(0); } -x_184 = lean_unsigned_to_nat(0u); -x_185 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_184, x_179); -x_186 = lean_nat_dec_eq(x_185, x_1); -lean_dec(x_185); -if (x_186 == 0) +} +block_31: { -lean_object* x_187; uint8_t x_188; lean_object* x_189; lean_object* x_190; -lean_dec(x_183); -lean_dec(x_180); -lean_dec(x_178); -x_187 = lean_box(0); -x_188 = lean_unbox(x_181); -lean_dec(x_181); -lean_inc(x_6); -x_189 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_188, x_182, x_179, x_6, x_187, x_8, x_9, x_10, x_11, x_177); -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -if (lean_obj_tag(x_190) == 0) +if (lean_obj_tag(x_25) == 0) { -lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +lean_object* x_27; lean_object* x_28; +lean_dec(x_23); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(x_1, x_11, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +lean_dec(x_27); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_dec(x_11); -lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_191 = lean_ctor_get(x_189, 1); -lean_inc(x_191); -if (lean_is_exclusive(x_189)) { - lean_ctor_release(x_189, 0); - lean_ctor_release(x_189, 1); - x_192 = x_189; +lean_dec(x_1); +x_29 = lean_ctor_get(x_25, 0); +lean_inc(x_29); +lean_dec(x_25); +if (lean_is_scalar(x_23)) { + x_30 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_189); - x_192 = lean_box(0); + x_30 = x_23; } -x_193 = lean_ctor_get(x_190, 0); -lean_inc(x_193); -lean_dec(x_190); -if (lean_is_scalar(x_192)) { - x_194 = lean_alloc_ctor(0, 2, 0); -} else { - x_194 = x_192; +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_26); +return x_30; } -lean_ctor_set(x_194, 0, x_193); -lean_ctor_set(x_194, 1, x_191); -return x_194; +} +block_54: +{ +uint8_t x_34; +x_34 = l_Lean_Exception_isInterrupt(x_32); +if (x_34 == 0) +{ +uint8_t x_35; +x_35 = l_Lean_Exception_isRuntime(x_32); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; uint8_t x_38; +lean_dec(x_32); +lean_dec(x_19); +x_36 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_33); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_unbox(x_37); +lean_dec(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_box(0); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_21, x_40, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_39); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_25 = x_42; +x_26 = x_43; +goto block_31; } else { -lean_object* x_195; lean_object* x_196; -x_195 = lean_ctor_get(x_189, 1); -lean_inc(x_195); -lean_dec(x_189); -x_196 = lean_ctor_get(x_190, 0); -lean_inc(x_196); -lean_dec(x_190); -x_7 = x_196; -x_12 = x_195; -goto _start; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_44 = lean_ctor_get(x_36, 1); +lean_inc(x_44); +lean_dec(x_36); +x_45 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2; +x_46 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2(x_24, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_44); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_21, x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_48); +lean_dec(x_47); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_25 = x_50; +x_26 = x_51; +goto block_31; } } else { -lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_object* x_52; +lean_dec(x_23); +lean_dec(x_21); lean_dec(x_11); -lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -if (lean_is_scalar(x_183)) { - x_198 = lean_alloc_ctor(0, 2, 0); -} else { - x_198 = x_183; -} -lean_ctor_set(x_198, 0, x_181); -lean_ctor_set(x_198, 1, x_182); -if (lean_is_scalar(x_180)) { - x_199 = lean_alloc_ctor(0, 2, 0); -} else { - x_199 = x_180; -} -lean_ctor_set(x_199, 0, x_179); -lean_ctor_set(x_199, 1, x_198); -x_200 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_201 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_199); -if (lean_is_scalar(x_178)) { - x_202 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_1); +if (lean_is_scalar(x_19)) { + x_52 = lean_alloc_ctor(1, 2, 0); } else { - x_202 = x_178; + x_52 = x_19; + lean_ctor_set_tag(x_52, 1); } -lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_177); -return x_202; +lean_ctor_set(x_52, 0, x_32); +lean_ctor_set(x_52, 1, x_33); +return x_52; } } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_object* x_53; +lean_dec(x_23); +lean_dec(x_21); lean_dec(x_11); -lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_203 = lean_ctor_get(x_174, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_174, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_205 = x_174; -} else { - lean_dec_ref(x_174); - x_205 = lean_box(0); -} -if (lean_is_scalar(x_205)) { - x_206 = lean_alloc_ctor(1, 2, 0); +lean_dec(x_1); +if (lean_is_scalar(x_19)) { + x_53 = lean_alloc_ctor(1, 2, 0); } else { - x_206 = x_205; + x_53 = x_19; + lean_ctor_set_tag(x_53, 1); } -lean_ctor_set(x_206, 0, x_203); -lean_ctor_set(x_206, 1, x_204); -return x_206; +lean_ctor_set(x_53, 0, x_32); +lean_ctor_set(x_53, 1, x_33); +return x_53; } } } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +case 7: { -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_6, 1); -lean_inc(x_12); -lean_dec(x_6); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_94; lean_object* x_95; uint8_t x_96; +x_94 = lean_ctor_get(x_1, 1); +lean_inc(x_94); +x_95 = lean_ctor_get(x_1, 2); +lean_inc(x_95); +x_96 = l_Lean_Expr_isArrow(x_1); +lean_dec(x_1); +if (x_96 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_12, 1); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +lean_object* x_97; lean_object* x_98; +lean_dec(x_95); +lean_dec(x_94); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_97 = lean_box(0); +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_10); +return x_98; +} +else { -lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; -x_16 = lean_ctor_get(x_14, 0); -lean_dec(x_16); -x_17 = lean_box(0); -x_18 = 0; -x_19 = lean_box(x_18); -lean_ctor_set(x_14, 0, x_19); -x_20 = lean_array_size(x_2); -x_21 = 0; -lean_inc(x_10); +lean_object* x_99; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_4); -lean_inc(x_3); -x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_17, x_2, x_20, x_21, x_12, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -x_25 = !lean_is_exclusive(x_22); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_22, 1); -x_27 = lean_ctor_get(x_22, 0); -lean_dec(x_27); -x_28 = !lean_is_exclusive(x_23); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = lean_ctor_get(x_23, 0); -x_30 = lean_ctor_get(x_23, 1); -lean_dec(x_30); -x_31 = !lean_is_exclusive(x_24); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_32 = lean_ctor_get(x_24, 0); -x_33 = lean_ctor_get(x_24, 1); -x_34 = lean_unsigned_to_nat(0u); -x_35 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_34, x_29); -x_36 = lean_nat_dec_eq(x_35, x_1); -lean_dec(x_35); -if (x_36 == 0) +lean_inc(x_6); +lean_inc(x_94); +x_99 = l_Lean_Meta_isProp(x_94, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_99) == 0) { -lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; -lean_free_object(x_24); -lean_free_object(x_23); -lean_free_object(x_22); -x_37 = lean_box(0); -x_38 = lean_unbox(x_32); -lean_dec(x_32); -lean_inc(x_5); -x_39 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_38, x_33, x_29, x_5, x_37, x_7, x_8, x_9, x_10, x_26); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 0) +lean_object* x_100; uint8_t x_101; +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_unbox(x_100); +lean_dec(x_100); +if (x_101 == 0) { -uint8_t x_41; -lean_dec(x_10); +uint8_t x_102; +lean_dec(x_95); +lean_dec(x_94); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_41 = !lean_is_exclusive(x_39); -if (x_41 == 0) +x_102 = !lean_is_exclusive(x_99); +if (x_102 == 0) { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_39, 0); -lean_dec(x_42); -x_43 = lean_ctor_get(x_40, 0); -lean_inc(x_43); -lean_dec(x_40); -lean_ctor_set(x_39, 0, x_43); -return x_39; +lean_object* x_103; lean_object* x_104; +x_103 = lean_ctor_get(x_99, 0); +lean_dec(x_103); +x_104 = lean_box(0); +lean_ctor_set(x_99, 0, x_104); +return x_99; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_39, 1); -lean_inc(x_44); -lean_dec(x_39); -x_45 = lean_ctor_get(x_40, 0); -lean_inc(x_45); -lean_dec(x_40); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_44); -return x_46; +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_99, 1); +lean_inc(x_105); +lean_dec(x_99); +x_106 = lean_box(0); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_105); +return x_107; } } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_39, 1); -lean_inc(x_47); -lean_dec(x_39); -x_48 = lean_ctor_get(x_40, 0); -lean_inc(x_48); -lean_dec(x_40); -x_6 = x_48; -x_11 = x_47; -goto _start; -} -} -else +lean_object* x_108; lean_object* x_109; +x_108 = lean_ctor_get(x_99, 1); +lean_inc(x_108); +lean_dec(x_99); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_95); +x_109 = l_Lean_Meta_isProp(x_95, x_6, x_7, x_8, x_9, x_108); +if (lean_obj_tag(x_109) == 0) { -lean_object* x_50; lean_object* x_51; -lean_dec(x_10); +lean_object* x_110; uint8_t x_111; +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_unbox(x_110); +lean_dec(x_110); +if (x_111 == 0) +{ +uint8_t x_112; +lean_dec(x_95); +lean_dec(x_94); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_50 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_23); -lean_ctor_set(x_22, 0, x_51); -return x_22; -} +x_112 = !lean_is_exclusive(x_109); +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_109, 0); +lean_dec(x_113); +x_114 = lean_box(0); +lean_ctor_set(x_109, 0, x_114); +return x_109; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_52 = lean_ctor_get(x_24, 0); -x_53 = lean_ctor_get(x_24, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_24); -x_54 = lean_unsigned_to_nat(0u); -x_55 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_54, x_29); -x_56 = lean_nat_dec_eq(x_55, x_1); -lean_dec(x_55); -if (x_56 == 0) +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_109, 1); +lean_inc(x_115); +lean_dec(x_109); +x_116 = lean_box(0); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_115); +return x_117; +} +} +else { -lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; -lean_free_object(x_23); -lean_free_object(x_22); -x_57 = lean_box(0); -x_58 = lean_unbox(x_52); -lean_dec(x_52); +lean_object* x_118; lean_object* x_119; +x_118 = lean_ctor_get(x_109, 1); +lean_inc(x_118); +lean_dec(x_109); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -x_59 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_58, x_53, x_29, x_5, x_57, x_7, x_8, x_9, x_10, x_26); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -if (lean_obj_tag(x_60) == 0) +lean_inc(x_4); +lean_inc(x_3); +x_119 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_94, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_118); +if (lean_obj_tag(x_119) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -lean_dec(x_10); +lean_object* x_120; lean_object* x_121; +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +lean_dec(x_119); +x_121 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_95, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_120); +return x_121; +} +else +{ +uint8_t x_122; +lean_dec(x_95); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_62 = x_59; -} else { - lean_dec_ref(x_59); - x_62 = lean_box(0); -} -x_63 = lean_ctor_get(x_60, 0); -lean_inc(x_63); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_64 = lean_alloc_ctor(0, 2, 0); -} else { - x_64 = x_62; -} -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_61); -return x_64; +x_122 = !lean_is_exclusive(x_119); +if (x_122 == 0) +{ +return x_119; } else { -lean_object* x_65; lean_object* x_66; -x_65 = lean_ctor_get(x_59, 1); -lean_inc(x_65); -lean_dec(x_59); -x_66 = lean_ctor_get(x_60, 0); -lean_inc(x_66); -lean_dec(x_60); -x_6 = x_66; -x_11 = x_65; -goto _start; +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_119, 0); +x_124 = lean_ctor_get(x_119, 1); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_119); +x_125 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; +} +} } } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_10); +uint8_t x_126; +lean_dec(x_95); +lean_dec(x_94); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_52); -lean_ctor_set(x_68, 1, x_53); -lean_ctor_set(x_23, 1, x_68); -x_69 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_23); -lean_ctor_set(x_22, 0, x_70); -return x_22; -} -} +x_126 = !lean_is_exclusive(x_109); +if (x_126 == 0) +{ +return x_109; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_71 = lean_ctor_get(x_23, 0); -lean_inc(x_71); -lean_dec(x_23); -x_72 = lean_ctor_get(x_24, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_24, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_74 = x_24; -} else { - lean_dec_ref(x_24); - x_74 = lean_box(0); +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_109, 0); +x_128 = lean_ctor_get(x_109, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_109); +x_129 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +return x_129; } -x_75 = lean_unsigned_to_nat(0u); -x_76 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_75, x_71); -x_77 = lean_nat_dec_eq(x_76, x_1); -lean_dec(x_76); -if (x_77 == 0) -{ -lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; -lean_dec(x_74); -lean_free_object(x_22); -x_78 = lean_box(0); -x_79 = lean_unbox(x_72); -lean_dec(x_72); -lean_inc(x_5); -x_80 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_79, x_73, x_71, x_5, x_78, x_7, x_8, x_9, x_10, x_26); -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -if (lean_obj_tag(x_81) == 0) +} +} +} +else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_10); +uint8_t x_130; +lean_dec(x_95); +lean_dec(x_94); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_83 = x_80; -} else { - lean_dec_ref(x_80); - x_83 = lean_box(0); -} -x_84 = lean_ctor_get(x_81, 0); -lean_inc(x_84); -lean_dec(x_81); -if (lean_is_scalar(x_83)) { - x_85 = lean_alloc_ctor(0, 2, 0); -} else { - x_85 = x_83; -} -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_82); -return x_85; +x_130 = !lean_is_exclusive(x_99); +if (x_130 == 0) +{ +return x_99; } else { -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_80, 1); -lean_inc(x_86); -lean_dec(x_80); -x_87 = lean_ctor_get(x_81, 0); -lean_inc(x_87); -lean_dec(x_81); -x_6 = x_87; -x_11 = x_86; -goto _start; +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_99, 0); +x_132 = lean_ctor_get(x_99, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_99); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +return x_133; } } -else +} +} +default: { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -lean_dec(x_10); +lean_object* x_134; lean_object* x_135; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -if (lean_is_scalar(x_74)) { - x_89 = lean_alloc_ctor(0, 2, 0); -} else { - x_89 = x_74; +lean_dec(x_1); +x_134 = lean_box(0); +x_135 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_10); +return x_135; +} } -lean_ctor_set(x_89, 0, x_72); -lean_ctor_set(x_89, 1, x_73); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_71); -lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_90); -lean_ctor_set(x_22, 0, x_92); -return x_22; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_st_ref_get(x_3, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get_uint8(x_11, sizeof(void*)*1); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_box(0); +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7(x_1, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +return x_15; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; -x_93 = lean_ctor_get(x_22, 1); -lean_inc(x_93); -lean_dec(x_22); -x_94 = lean_ctor_get(x_23, 0); -lean_inc(x_94); -if (lean_is_exclusive(x_23)) { - lean_ctor_release(x_23, 0); - lean_ctor_release(x_23, 1); - x_95 = x_23; -} else { - lean_dec_ref(x_23); - x_95 = lean_box(0); +uint8_t x_16; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_16 = !lean_is_exclusive(x_10); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_10, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_10, 0, x_18); +return x_10; } -x_96 = lean_ctor_get(x_24, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_24, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_98 = x_24; -} else { - lean_dec_ref(x_24); - x_98 = lean_box(0); +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_dec(x_10); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } -x_99 = lean_unsigned_to_nat(0u); -x_100 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_99, x_94); -x_101 = lean_nat_dec_eq(x_100, x_1); -lean_dec(x_100); -if (x_101 == 0) +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; -lean_dec(x_98); -lean_dec(x_95); -x_102 = lean_box(0); -x_103 = lean_unbox(x_96); -lean_dec(x_96); -lean_inc(x_5); -x_104 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_103, x_97, x_94, x_5, x_102, x_7, x_8, x_9, x_10, x_93); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -if (lean_obj_tag(x_105) == 0) +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_16 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___spec__1(x_1, x_2, x_3, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_10); +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_106 = lean_ctor_get(x_104, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_104)) { - lean_ctor_release(x_104, 0); - lean_ctor_release(x_104, 1); - x_107 = x_104; -} else { - lean_dec_ref(x_104); - x_107 = lean_box(0); -} -x_108 = lean_ctor_get(x_105, 0); -lean_inc(x_108); -lean_dec(x_105); -if (lean_is_scalar(x_107)) { - x_109 = lean_alloc_ctor(0, 2, 0); -} else { - x_109 = x_107; -} -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_106); -return x_109; -} -else -{ -lean_object* x_110; lean_object* x_111; -x_110 = lean_ctor_get(x_104, 1); -lean_inc(x_110); -lean_dec(x_104); -x_111 = lean_ctor_get(x_105, 0); -lean_inc(x_111); -lean_dec(x_105); -x_6 = x_111; -x_11 = x_110; -goto _start; +lean_dec(x_2); +return x_11; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -if (lean_is_scalar(x_98)) { - x_113 = lean_alloc_ctor(0, 2, 0); -} else { - x_113 = x_98; -} -lean_ctor_set(x_113, 0, x_96); -lean_ctor_set(x_113, 1, x_97); -if (lean_is_scalar(x_95)) { - x_114 = lean_alloc_ctor(0, 2, 0); -} else { - x_114 = x_95; +return x_12; } -lean_ctor_set(x_114, 0, x_94); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_116 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_114); -x_117 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_93); -return x_117; } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_6); +lean_dec(x_4); +return x_13; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -uint8_t x_118; -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_118 = !lean_is_exclusive(x_22); -if (x_118 == 0) +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -return x_22; +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +lean_dec(x_1); +return x_13; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_22, 0); -x_120 = lean_ctor_get(x_22, 1); -lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_22); -x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -return x_121; +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_2); +return x_11; } } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_6, x_5); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_7); +lean_ctor_set(x_17, 1, x_15); +return x_17; } else { -lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; -x_122 = lean_ctor_get(x_14, 1); -lean_inc(x_122); -lean_dec(x_14); -x_123 = lean_box(0); -x_124 = 0; -x_125 = lean_box(x_124); -x_126 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_122); -lean_ctor_set(x_12, 1, x_126); -x_127 = lean_array_size(x_2); -x_128 = 0; +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_28; lean_object* x_29; +lean_dec(x_7); +x_18 = lean_array_uget(x_4, x_6); +x_28 = 1; +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_29 = l_Lean_Meta_Grind_preprocessPattern(x_18, x_28, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_4); +x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect(x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +x_34 = lean_st_ref_get(x_9, x_33); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get_uint8(x_35, sizeof(void*)*1); +lean_dec(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); lean_inc(x_3); -x_129 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_123, x_2, x_127, x_128, x_12, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_129) == 0) +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_3); +x_19 = x_38; +x_20 = x_37; +goto block_27; +} +else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; -x_130 = lean_ctor_get(x_129, 0); -lean_inc(x_130); -x_131 = lean_ctor_get(x_130, 1); -lean_inc(x_131); -x_132 = lean_ctor_get(x_129, 1); -lean_inc(x_132); -if (lean_is_exclusive(x_129)) { - lean_ctor_release(x_129, 0); - lean_ctor_release(x_129, 1); - x_133 = x_129; -} else { - lean_dec_ref(x_129); - x_133 = lean_box(0); +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_34, 1); +lean_inc(x_39); +lean_dec(x_34); +x_40 = lean_st_ref_get(x_9, x_39); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_42 = lean_ctor_get(x_40, 0); +x_43 = lean_ctor_get(x_40, 1); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_array_to_list(x_44); +x_46 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_47, 0, x_46); +x_48 = lean_box(0); +lean_ctor_set(x_40, 1, x_48); +lean_ctor_set(x_40, 0, x_47); +x_49 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_49, 0, x_40); +x_19 = x_49; +x_20 = x_43; +goto block_27; } -x_134 = lean_ctor_get(x_130, 0); -lean_inc(x_134); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_135 = x_130; -} else { - lean_dec_ref(x_130); - x_135 = lean_box(0); +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_50 = lean_ctor_get(x_40, 0); +x_51 = lean_ctor_get(x_40, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_40); +x_52 = lean_ctor_get(x_50, 0); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_array_to_list(x_52); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_53); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_58, 0, x_57); +x_19 = x_58; +x_20 = x_51; +goto block_27; } -x_136 = lean_ctor_get(x_131, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_131, 1); -lean_inc(x_137); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_138 = x_131; -} else { - lean_dec_ref(x_131); - x_138 = lean_box(0); } -x_139 = lean_unsigned_to_nat(0u); -x_140 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_139, x_134); -x_141 = lean_nat_dec_eq(x_140, x_1); -lean_dec(x_140); -if (x_141 == 0) -{ -lean_object* x_142; uint8_t x_143; lean_object* x_144; lean_object* x_145; -lean_dec(x_138); -lean_dec(x_135); -lean_dec(x_133); -x_142 = lean_box(0); -x_143 = lean_unbox(x_136); -lean_dec(x_136); -lean_inc(x_5); -x_144 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_143, x_137, x_134, x_5, x_142, x_7, x_8, x_9, x_10, x_132); -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -if (lean_obj_tag(x_145) == 0) +} +else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +uint8_t x_59; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -x_146 = lean_ctor_get(x_144, 1); -lean_inc(x_146); -if (lean_is_exclusive(x_144)) { - lean_ctor_release(x_144, 0); - lean_ctor_release(x_144, 1); - x_147 = x_144; -} else { - lean_dec_ref(x_144); - x_147 = lean_box(0); -} -x_148 = lean_ctor_get(x_145, 0); -lean_inc(x_148); -lean_dec(x_145); -if (lean_is_scalar(x_147)) { - x_149 = lean_alloc_ctor(0, 2, 0); -} else { - x_149 = x_147; -} -lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_146); -return x_149; +x_59 = !lean_is_exclusive(x_32); +if (x_59 == 0) +{ +return x_32; } else { -lean_object* x_150; lean_object* x_151; -x_150 = lean_ctor_get(x_144, 1); -lean_inc(x_150); -lean_dec(x_144); -x_151 = lean_ctor_get(x_145, 0); -lean_inc(x_151); -lean_dec(x_145); -x_6 = x_151; -x_11 = x_150; -goto _start; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_32, 0); +x_61 = lean_ctor_get(x_32, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_32); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} } } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +uint8_t x_63; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -if (lean_is_scalar(x_138)) { - x_153 = lean_alloc_ctor(0, 2, 0); -} else { - x_153 = x_138; -} -lean_ctor_set(x_153, 0, x_136); -lean_ctor_set(x_153, 1, x_137); -if (lean_is_scalar(x_135)) { - x_154 = lean_alloc_ctor(0, 2, 0); -} else { - x_154 = x_135; -} -lean_ctor_set(x_154, 0, x_134); -lean_ctor_set(x_154, 1, x_153); -x_155 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_156 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_154); -if (lean_is_scalar(x_133)) { - x_157 = lean_alloc_ctor(0, 2, 0); -} else { - x_157 = x_133; +x_63 = !lean_is_exclusive(x_29); +if (x_63 == 0) +{ +return x_29; } -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_132); -return x_157; +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_29, 0); +x_65 = lean_ctor_get(x_29, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_29); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } -else +block_27: { -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -x_158 = lean_ctor_get(x_129, 0); -lean_inc(x_158); -x_159 = lean_ctor_get(x_129, 1); -lean_inc(x_159); -if (lean_is_exclusive(x_129)) { - lean_ctor_release(x_129, 0); - lean_ctor_release(x_129, 1); - x_160 = x_129; -} else { - lean_dec_ref(x_129); - x_160 = lean_box(0); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } -if (lean_is_scalar(x_160)) { - x_161 = lean_alloc_ctor(1, 2, 0); -} else { - x_161 = x_160; +else +{ +lean_object* x_23; size_t x_24; size_t x_25; +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +lean_dec(x_19); +x_24 = 1; +x_25 = lean_usize_add(x_6, x_24); +x_6 = x_25; +x_7 = x_23; +x_15 = x_20; +goto _start; } -lean_ctor_set(x_161, 0, x_158); -lean_ctor_set(x_161, 1, x_159); -return x_161; } } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; size_t x_171; size_t x_172; lean_object* x_173; -x_162 = lean_ctor_get(x_12, 1); -x_163 = lean_ctor_get(x_12, 0); -lean_inc(x_162); -lean_inc(x_163); -lean_dec(x_12); -x_164 = lean_ctor_get(x_162, 1); -lean_inc(x_164); -if (lean_is_exclusive(x_162)) { - lean_ctor_release(x_162, 0); - lean_ctor_release(x_162, 1); - x_165 = x_162; -} else { - lean_dec_ref(x_162); - x_165 = lean_box(0); +lean_object* x_11; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_10); +return x_11; } -x_166 = lean_box(0); -x_167 = 0; -x_168 = lean_box(x_167); -if (lean_is_scalar(x_165)) { - x_169 = lean_alloc_ctor(0, 2, 0); -} else { - x_169 = x_165; } -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_164); -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_163); -lean_ctor_set(x_170, 1, x_169); -x_171 = lean_array_size(x_2); -x_172 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_4); -lean_inc(x_3); -x_173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_2, x_3, x_4, x_166, x_2, x_171, x_172, x_170, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_173) == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1() { +_start: { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; uint8_t x_185; -x_174 = lean_ctor_get(x_173, 0); -lean_inc(x_174); -x_175 = lean_ctor_get(x_174, 1); -lean_inc(x_175); -x_176 = lean_ctor_get(x_173, 1); -lean_inc(x_176); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_177 = x_173; -} else { - lean_dec_ref(x_173); - x_177 = lean_box(0); +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_2 = 0; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; } -x_178 = lean_ctor_get(x_174, 0); -lean_inc(x_178); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_179 = x_174; +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_9 = lean_box(0); +x_10 = lean_box(0); +x_11 = lean_array_size(x_3); +x_12 = 0; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +x_14 = l_Lean_Meta_Grind_NormalizePattern_main___closed__5; +x_15 = lean_st_mk_ref(x_14, x_8); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + lean_ctor_release(x_15, 1); + x_18 = x_15; } else { - lean_dec_ref(x_174); - x_179 = lean_box(0); + lean_dec_ref(x_15); + x_18 = lean_box(0); +} +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1; +x_50 = lean_st_mk_ref(x_49, x_17); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_canBeSynthesized___spec__3___closed__1; +lean_inc(x_16); +lean_inc(x_51); +x_54 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1(x_3, x_9, x_53, x_3, x_11, x_12, x_53, x_13, x_51, x_16, x_4, x_5, x_6, x_7, x_52); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +lean_dec(x_55); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +lean_dec(x_54); +x_58 = lean_st_ref_get(x_51, x_57); +lean_dec(x_51); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_19 = x_10; +x_20 = x_59; +goto block_48; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_54, 1); +lean_inc(x_60); +lean_dec(x_54); +x_61 = lean_ctor_get(x_56, 0); +lean_inc(x_61); +lean_dec(x_56); +x_62 = lean_st_ref_get(x_51, x_60); +lean_dec(x_51); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_19 = x_61; +x_20 = x_63; +goto block_48; +} +} +else +{ +uint8_t x_64; +lean_dec(x_51); +lean_dec(x_18); +lean_dec(x_16); +x_64 = !lean_is_exclusive(x_54); +if (x_64 == 0) +{ +return x_54; } -x_180 = lean_ctor_get(x_175, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_175, 1); -lean_inc(x_181); -if (lean_is_exclusive(x_175)) { - lean_ctor_release(x_175, 0); - lean_ctor_release(x_175, 1); - x_182 = x_175; -} else { - lean_dec_ref(x_175); - x_182 = lean_box(0); +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_54, 0); +x_66 = lean_ctor_get(x_54, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_54); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; } -x_183 = lean_unsigned_to_nat(0u); -x_184 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_183, x_178); -x_185 = lean_nat_dec_eq(x_184, x_1); -lean_dec(x_184); -if (x_185 == 0) +} +block_48: { -lean_object* x_186; uint8_t x_187; lean_object* x_188; lean_object* x_189; -lean_dec(x_182); -lean_dec(x_179); -lean_dec(x_177); -x_186 = lean_box(0); -x_187 = lean_unbox(x_180); -lean_dec(x_180); -lean_inc(x_5); -x_188 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_187, x_181, x_178, x_5, x_186, x_7, x_8, x_9, x_10, x_176); -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -if (lean_obj_tag(x_189) == 0) +lean_object* x_21; +x_21 = lean_st_ref_get(x_16, x_20); +lean_dec(x_16); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_190 = lean_ctor_get(x_188, 1); -lean_inc(x_190); -if (lean_is_exclusive(x_188)) { - lean_ctor_release(x_188, 0); - lean_ctor_release(x_188, 1); - x_191 = x_188; -} else { - lean_dec_ref(x_188); - x_191 = lean_box(0); +uint8_t x_22; +lean_dec(x_18); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_10); +return x_21; } -x_192 = lean_ctor_get(x_189, 0); -lean_inc(x_192); -lean_dec(x_189); -if (lean_is_scalar(x_191)) { - x_193 = lean_alloc_ctor(0, 2, 0); +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_10); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_21); +if (x_26 == 0) +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_19); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_21, 0); +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_array_to_list(x_30); +if (lean_is_scalar(x_18)) { + x_32 = lean_alloc_ctor(0, 2, 0); } else { - x_193 = x_191; + x_32 = x_18; } -lean_ctor_set(x_193, 0, x_192); -lean_ctor_set(x_193, 1, x_190); -return x_193; +lean_ctor_set(x_32, 0, x_29); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_19, 0, x_32); +lean_ctor_set(x_21, 0, x_19); +return x_21; } else { -lean_object* x_194; lean_object* x_195; -x_194 = lean_ctor_get(x_188, 1); -lean_inc(x_194); -lean_dec(x_188); -x_195 = lean_ctor_get(x_189, 0); -lean_inc(x_195); -lean_dec(x_189); -x_6 = x_195; -x_11 = x_194; -goto _start; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_19, 0); +lean_inc(x_34); +lean_dec(x_19); +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_array_to_list(x_35); +if (lean_is_scalar(x_18)) { + x_37 = lean_alloc_ctor(0, 2, 0); +} else { + x_37 = x_18; +} +lean_ctor_set(x_37, 0, x_34); +lean_ctor_set(x_37, 1, x_36); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_21, 0, x_38); +return x_21; } } else { -lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -if (lean_is_scalar(x_182)) { - x_197 = lean_alloc_ctor(0, 2, 0); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_39 = lean_ctor_get(x_21, 0); +x_40 = lean_ctor_get(x_21, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_21); +x_41 = lean_ctor_get(x_19, 0); +lean_inc(x_41); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + x_42 = x_19; } else { - x_197 = x_182; + lean_dec_ref(x_19); + x_42 = lean_box(0); } -lean_ctor_set(x_197, 0, x_180); -lean_ctor_set(x_197, 1, x_181); -if (lean_is_scalar(x_179)) { - x_198 = lean_alloc_ctor(0, 2, 0); +x_43 = lean_ctor_get(x_39, 0); +lean_inc(x_43); +lean_dec(x_39); +x_44 = lean_array_to_list(x_43); +if (lean_is_scalar(x_18)) { + x_45 = lean_alloc_ctor(0, 2, 0); } else { - x_198 = x_179; + x_45 = x_18; } -lean_ctor_set(x_198, 0, x_178); -lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___closed__1; -x_200 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_200, 0, x_199); -lean_ctor_set(x_200, 1, x_198); -if (lean_is_scalar(x_177)) { - x_201 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_41); +lean_ctor_set(x_45, 1, x_44); +if (lean_is_scalar(x_42)) { + x_46 = lean_alloc_ctor(1, 1, 0); } else { - x_201 = x_177; + x_46 = x_42; } -lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_176); -return x_201; +lean_ctor_set(x_46, 0, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_40); +return x_47; } } -else +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: { -lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -lean_dec(x_10); +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_17 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___spec__1(x_1, x_2, x_3, x_4, x_16, x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_202 = lean_ctor_get(x_173, 0); -lean_inc(x_202); -x_203 = lean_ctor_get(x_173, 1); -lean_inc(x_203); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_204 = x_173; -} else { - lean_dec_ref(x_173); - x_204 = lean_box(0); +lean_dec(x_2); +return x_11; } -if (lean_is_scalar(x_204)) { - x_205 = lean_alloc_ctor(1, 2, 0); -} else { - x_205 = x_204; } -lean_ctor_set(x_205, 0, x_202); -lean_ctor_set(x_205, 1, x_203); -return x_205; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +return x_9; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +lean_ctor_set(x_13, 2, x_3); +lean_ctor_set(x_13, 3, x_4); +lean_ctor_set(x_13, 4, x_5); +lean_ctor_set(x_13, 5, x_6); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_12); +return x_15; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_5, 1); -x_16 = lean_nat_dec_lt(x_7, x_15); -if (x_16 == 0) -{ -lean_object* x_17; -lean_dec(x_7); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_14); -return x_17; -} -else +lean_object* x_11; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_3); +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_array_fget(x_1, x_7); -x_19 = l_Lean_Expr_fvarId_x21(x_18); -lean_dec(x_18); -x_20 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_3, x_19); -lean_dec(x_19); -if (lean_obj_tag(x_20) == 0) +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_inc(x_7); -x_21 = lean_array_push(x_6, x_7); -x_22 = lean_ctor_get(x_5, 2); -x_23 = lean_nat_add(x_7, x_22); +uint8_t x_13; +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_6 = x_21; -x_7 = x_23; -x_8 = lean_box(0); -x_9 = lean_box(0); -goto _start; +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_11, 0, x_15); +return x_11; } else { -lean_object* x_25; lean_object* x_26; -lean_dec(x_20); -x_25 = lean_ctor_get(x_5, 2); -x_26 = lean_nat_add(x_7, x_25); -lean_dec(x_7); -x_7 = x_26; -x_8 = lean_box(0); -x_9 = lean_box(0); -goto _start; -} -} +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +else { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_4, 1); -x_15 = lean_nat_dec_lt(x_6, x_14); -if (x_15 == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_12, 0); +lean_inc(x_19); +lean_dec(x_12); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_dec(x_11); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) { -lean_object* x_16; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_22 = lean_ctor_get(x_19, 0); +x_23 = lean_ctor_get(x_19, 1); +x_24 = lean_array_get_size(x_4); +lean_dec(x_4); +x_25 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_25, x_6, x_7, x_8, x_9, x_20); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_unbox(x_27); +lean_dec(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_free_object(x_19); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_dec(x_26); +x_30 = lean_box(0); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_24, x_22, x_23, x_1, x_30, x_6, x_7, x_8, x_9, x_29); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_5); -lean_ctor_set(x_16, 1, x_13); -return x_16; +return x_31; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_array_fget(x_1, x_6); -x_18 = l_Lean_Expr_fvarId_x21(x_17); -lean_dec(x_17); -x_19 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_2, x_18); -lean_dec(x_18); -if (lean_obj_tag(x_19) == 0) +uint8_t x_32; +x_32 = !lean_is_exclusive(x_26); +if (x_32 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_inc(x_6); -x_20 = lean_array_push(x_5, x_6); -x_21 = lean_ctor_get(x_4, 2); -x_22 = lean_nat_add(x_6, x_21); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_26, 1); +x_34 = lean_ctor_get(x_26, 0); +lean_dec(x_34); +lean_inc(x_1); +x_35 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_1, x_6, x_7, x_8, x_9, x_33); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_Lean_Meta_Grind_ppPattern___closed__4; +lean_ctor_set_tag(x_26, 7); +lean_ctor_set(x_26, 1, x_36); +lean_ctor_set(x_26, 0, x_38); +x_39 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_39); +lean_ctor_set(x_19, 0, x_26); +x_40 = lean_box(0); +lean_inc(x_22); +x_41 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_22, x_40); +x_42 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_41, x_40); +x_43 = l_Lean_MessageData_ofList(x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_19); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_38); +x_46 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_25, x_45, x_6, x_7, x_8, x_9, x_37); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_24, x_22, x_23, x_1, x_47, x_6, x_7, x_8, x_9, x_48); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_5 = x_20; -x_6 = x_22; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_dec(x_47); +return x_49; } else { -lean_object* x_24; lean_object* x_25; -lean_dec(x_19); -x_24 = lean_ctor_get(x_4, 2); -x_25 = lean_nat_add(x_6, x_24); +uint8_t x_50; +lean_free_object(x_26); +lean_dec(x_24); +lean_free_object(x_19); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_6 = x_25; -x_7 = lean_box(0); -x_8 = lean_box(0); -goto _start; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_50 = !lean_is_exclusive(x_35); +if (x_50 == 0) +{ +return x_35; } +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_35, 0); +x_52 = lean_ctor_get(x_35, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_35); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: +else { -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_2, x_3); -if (x_5 == 0) +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_26, 1); +lean_inc(x_54); +lean_dec(x_26); +lean_inc(x_1); +x_55 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_1, x_6, x_7, x_8, x_9, x_54); +if (lean_obj_tag(x_55) == 0) { -size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; -x_6 = 1; -x_7 = lean_usize_sub(x_2, x_6); -x_8 = lean_array_uget(x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_4, x_8); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = l_Lean_Meta_Grind_ppPattern___closed__4; +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_56); +x_60 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_60); +lean_ctor_set(x_19, 0, x_59); +x_61 = lean_box(0); +lean_inc(x_22); +x_62 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_22, x_61); +x_63 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_62, x_61); +x_64 = l_Lean_MessageData_ofList(x_63); +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_19); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_58); +x_67 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_25, x_66, x_6, x_7, x_8, x_9, x_57); +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_24, x_22, x_23, x_1, x_68, x_6, x_7, x_8, x_9, x_69); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_4); -x_2 = x_7; -x_4 = x_9; -goto _start; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_68); +return x_70; } else { -return x_4; -} +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_24); +lean_free_object(x_19); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_71 = lean_ctor_get(x_55, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_55, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_73 = x_55; +} else { + lean_dec_ref(x_55); + x_73 = lean_box(0); } +if (lean_is_scalar(x_73)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_73; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_array_mk(x_1); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_2); -lean_ctor_set(x_14, 2, x_13); -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_3, x_4, x_14, x_14, x_11, x_12, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10); -lean_dec(x_14); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_array_to_list(x_17); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_15, 0, x_19); -return x_15; +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +return x_74; } -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_15, 0); -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_15); -x_22 = lean_array_to_list(x_20); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -return x_24; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_14 = lean_box(0); -x_15 = 0; -x_16 = lean_box(x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_17); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_14); -lean_ctor_set(x_19, 1, x_18); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_4); -x_20 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_2, x_3, x_4, x_5, x_14, x_19, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 0); -lean_inc(x_23); -lean_dec(x_21); -if (lean_obj_tag(x_23) == 0) +else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_ctor_get(x_22, 0); -lean_inc(x_25); -lean_dec(x_22); -x_26 = lean_box(0); -x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_4, x_6, x_3, x_25, x_26, x_9, x_10, x_11, x_12, x_24); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; +x_75 = lean_ctor_get(x_19, 0); +x_76 = lean_ctor_get(x_19, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_19); +x_77 = lean_array_get_size(x_4); +lean_dec(x_4); +x_78 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3; +x_79 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_78, x_6, x_7, x_8, x_9, x_20); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_unbox(x_80); +lean_dec(x_80); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_79, 1); +lean_inc(x_82); +lean_dec(x_79); +x_83 = lean_box(0); +x_84 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_77, x_75, x_76, x_1, x_83, x_6, x_7, x_8, x_9, x_82); lean_dec(x_9); -lean_dec(x_25); -return x_27; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_84; } else { -uint8_t x_28; -lean_dec(x_22); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_79, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_86 = x_79; +} else { + lean_dec_ref(x_79); + x_86 = lean_box(0); +} +lean_inc(x_1); +x_87 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_1, x_6, x_7, x_8, x_9, x_85); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_Meta_Grind_ppPattern___closed__4; +if (lean_is_scalar(x_86)) { + x_91 = lean_alloc_ctor(7, 2, 0); +} else { + x_91 = x_86; + lean_ctor_set_tag(x_91, 7); +} +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_88); +x_92 = l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5; +x_93 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_box(0); +lean_inc(x_75); +x_95 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_75, x_94); +x_96 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_95, x_94); +x_97 = l_Lean_MessageData_ofList(x_96); +x_98 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_98, 0, x_93); +lean_ctor_set(x_98, 1, x_97); +x_99 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_90); +x_100 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_78, x_99, x_6, x_7, x_8, x_9, x_89); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +lean_dec(x_100); +x_103 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_2, x_3, x_77, x_75, x_76, x_1, x_101, x_6, x_7, x_8, x_9, x_102); lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); -x_28 = !lean_is_exclusive(x_20); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_20, 0); -lean_dec(x_29); -x_30 = lean_ctor_get(x_23, 0); -lean_inc(x_30); -lean_dec(x_23); -lean_ctor_set(x_20, 0, x_30); -return x_20; +lean_dec(x_101); +return x_103; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_20, 1); -lean_inc(x_31); -lean_dec(x_20); -x_32 = lean_ctor_get(x_23, 0); -lean_inc(x_32); -lean_dec(x_23); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -return x_33; +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_86); +lean_dec(x_77); +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_104 = lean_ctor_get(x_87, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_87, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_106 = x_87; +} else { + lean_dec_ref(x_87); + x_106 = lean_box(0); +} +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(1, 2, 0); +} else { + x_107 = x_106; +} +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_105); +return x_107; +} +} } } } else { -uint8_t x_34; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_108; lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_34 = !lean_is_exclusive(x_20); -if (x_34 == 0) +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_108 = !lean_is_exclusive(x_11); +if (x_108 == 0) { -return x_20; +return x_11; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_20, 0); -x_36 = lean_ctor_get(x_20, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_20); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_11, 0); +x_110 = lean_ctor_get(x_11, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_11); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("numParams == xs.size\n ", 25, 25); -return x_1; +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_13; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__1; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +return x_11; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3() { +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; +x_8 = lean_panic_fn(x_7, x_1); +x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_1, x_2, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.checkCoverage", 77, 77); +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.EMatchTheorem.0.Lean.Meta.Grind.mkEMatchTheoremWithKind\?", 88, 88); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__3; -x_3 = lean_unsigned_to_nat(263u); -x_4 = lean_unsigned_to_nat(4u); -x_5 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1; +x_3 = lean_unsigned_to_nat(616u); +x_4 = lean_unsigned_to_nat(13u); +x_5 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3() { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_array_get_size(x_3); -x_11 = lean_nat_dec_eq(x_1, x_10); -if (x_11 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid `grind` forward theorem, theorem `", 42, 42); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4() { +_start: { -lean_object* x_12; lean_object* x_13; -lean_dec(x_10); -lean_dec(x_3); -x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___closed__4; -x_13 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1(x_12, x_5, x_6, x_7, x_8, x_9); -return x_13; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` does not have propositional hypotheses", 40, 40); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = lean_box(x_1); +switch (lean_obj_tag(x_12)) { +case 3: +{ +lean_object* x_13; +lean_dec(x_6); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(x_5, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Array_isEmpty___rarg(x_14); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_15); +lean_dec(x_14); +return x_17; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_48; uint8_t x_49; -x_14 = lean_box(0); -x_15 = lean_ctor_get(x_2, 1); -x_16 = lean_array_get_size(x_15); -lean_inc(x_3); -x_17 = lean_array_to_list(x_3); -x_18 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__3(x_17, x_14); -x_19 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_18); -x_20 = lean_box(0); -x_48 = lean_unsigned_to_nat(0u); -x_49 = lean_nat_dec_lt(x_48, x_16); -if (x_49 == 0) +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_18 = l_Lean_Meta_Grind_Origin_pp___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__3(x_2, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_18) == 0) { -lean_dec(x_16); -x_21 = x_14; -goto block_47; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4; +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +x_23 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_24, x_7, x_8, x_9, x_10, x_20); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +return x_25; } else { -size_t x_50; size_t x_51; lean_object* x_52; -x_50 = lean_usize_of_nat(x_16); -lean_dec(x_16); -x_51 = 0; -x_52 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_15, x_50, x_51, x_14); -x_21 = x_52; -goto block_47; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_25); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } -block_47: -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_3, x_10, x_21, x_14); -lean_inc(x_22); -x_23 = l_Lean_RBTree_ofList___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__4(x_22); -lean_inc(x_5); -lean_inc(x_23); -lean_inc(x_22); -x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_14, x_19, x_20, x_22, x_22, x_22, x_23, lean_box(0), x_5, x_6, x_7, x_8, x_9); -lean_dec(x_22); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_26 = lean_ctor_get(x_24, 0); -x_27 = lean_ctor_get(x_24, 1); -x_28 = lean_unsigned_to_nat(0u); -x_29 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_28, x_26); -x_30 = lean_nat_dec_eq(x_29, x_1); -lean_dec(x_29); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; -lean_free_object(x_24); -x_31 = lean_box(0); -x_32 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_26, x_31, x_5, x_6, x_7, x_8, x_27); -lean_dec(x_3); -return x_32; } else { -lean_object* x_33; -lean_dec(x_26); -lean_dec(x_23); -lean_dec(x_19); +uint8_t x_30; lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_33 = lean_box(0); -lean_ctor_set(x_24, 0, x_33); -return x_24; -} +x_30 = !lean_is_exclusive(x_18); +if (x_30 == 0) +{ +return x_18; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_34 = lean_ctor_get(x_24, 0); -x_35 = lean_ctor_get(x_24, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_24); -x_36 = lean_unsigned_to_nat(0u); -x_37 = l_Lean_RBNode_fold___at_Lean_RBMap_size___spec__1___rarg(x_36, x_34); -x_38 = lean_nat_dec_eq(x_37, x_1); -lean_dec(x_37); -if (x_38 == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_box(0); -x_40 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_23, x_1, x_3, x_14, x_19, x_10, x_34, x_39, x_5, x_6, x_7, x_8, x_35); -lean_dec(x_3); -return x_40; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_18, 0); +x_32 = lean_ctor_get(x_18, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_18); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} } else { -lean_object* x_41; lean_object* x_42; -lean_dec(x_34); -lean_dec(x_23); -lean_dec(x_19); +uint8_t x_34; lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_41 = lean_box(0); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_35); -return x_42; +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_13); +if (x_34 == 0) +{ +return x_13; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_13, 0); +x_36 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_13); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} } } +case 4: +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_6); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_array_mk(x_39); +x_41 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_2, x_3, x_4, x_5, x_40, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_40); +return x_41; +} +case 5: +{ +lean_object* x_42; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_42 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getPropTypes(x_5, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_6); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_array_mk(x_46); +x_48 = l_Array_append___rarg(x_47, x_43); +lean_dec(x_43); +x_49 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_2, x_3, x_4, x_5, x_48, x_7, x_8, x_9, x_10, x_44); +lean_dec(x_48); +return x_49; } else { -uint8_t x_43; -lean_dec(x_23); -lean_dec(x_19); +uint8_t x_50; lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_43 = !lean_is_exclusive(x_24); -if (x_43 == 0) +lean_dec(x_2); +x_50 = !lean_is_exclusive(x_42); +if (x_50 == 0) { -return x_24; +return x_42; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_24, 0); -x_45 = lean_ctor_get(x_24, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_24); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} -} -} +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_42, 0); +x_52 = lean_ctor_get(x_42, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_42); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +default: { -lean_object* x_10; +lean_object* x_54; lean_object* x_55; +lean_dec(x_12); +lean_dec(x_6); +x_54 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2; +lean_inc(x_10); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_10 = lean_infer_type(x_1, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) +x_55 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___spec__1(x_54, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_55) == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -lean_inc(x_2); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_2); -x_14 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed), 9, 2); -lean_closure_set(x_14, 0, x_2); -lean_closure_set(x_14, 1, x_3); -x_15 = 0; -x_16 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_11, x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_12); -return x_16; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f_go(x_2, x_3, x_4, x_5, x_56, x_7, x_8, x_9, x_10, x_57); +lean_dec(x_56); +return x_58; } else { -uint8_t x_17; +uint8_t x_59; +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = !lean_is_exclusive(x_10); -if (x_17 == 0) +x_59 = !lean_is_exclusive(x_55); +if (x_59 == 0) { -return x_10; +return x_55; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_10, 0); -x_19 = lean_ctor_get(x_10, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_10); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_55, 0); +x_61 = lean_ctor_get(x_55, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_55); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_3, 0); +lean_object* x_11; lean_inc(x_9); -x_10 = lean_nat_dec_eq(x_9, x_2); -lean_dec(x_9); -if (x_10 == 0) +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_11, x_4, x_5, x_6, x_7, x_8); -return x_12; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_box(x_2); +x_15 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___boxed), 11, 4); +lean_closure_set(x_15, 0, x_14); +lean_closure_set(x_15, 1, x_3); +lean_closure_set(x_15, 2, x_4); +lean_closure_set(x_15, 3, x_1); +x_16 = 0; +x_17 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_12, x_15, x_16, x_6, x_7, x_8, x_9, x_13); +return x_17; } else { -lean_object* x_13; lean_object* x_14; +uint8_t x_18; +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_8); -return x_14; +x_18 = !lean_is_exclusive(x_11); +if (x_18 == 0) +{ +return x_11; } +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_11, 0); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_foldrM___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__2(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_6; -x_6 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__5(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_6; +uint8_t x_10; uint8_t x_11; +x_10 = 0; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_4, x_10); +if (x_11 == 0) +{ +uint8_t x_12; uint8_t x_13; +x_12 = 1; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_4, x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_box(0); +x_15 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3(x_3, x_4, x_1, x_2, x_14, x_5, x_6, x_7, x_8, x_9); +return x_15; } +else +{ +uint8_t x_16; lean_object* x_17; +x_16 = 0; +x_17 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore(x_1, x_2, x_3, x_16, x_16, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_17, 0, x_20); +return x_17; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_17); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_21); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__6___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__7(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} +uint8_t x_25; +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) +{ +return x_17; } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: +else { -lean_object* x_15; -x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_15; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_17, 0); +x_27 = lean_ctor_get(x_17, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__8___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_14; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} +uint8_t x_29; uint8_t x_30; lean_object* x_31; +x_29 = 0; +x_30 = 1; +x_31 = l_Lean_Meta_Grind_mkEMatchEqTheoremCore(x_1, x_2, x_3, x_29, x_30, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_31, 0, x_34); +return x_31; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__10___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__11(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_31, 0); +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_31); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_35); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12(x_1, x_2, x_3, x_7, x_8, x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} +uint8_t x_39; +x_39 = !lean_is_exclusive(x_31); +if (x_39 == 0) +{ +return x_31; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__12___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__13(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_31, 0); +x_41 = lean_ctor_get(x_31, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_31); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_16, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -return x_17; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -size_t x_14; size_t x_15; lean_object* x_16; -x_14 = lean_unbox_usize(x_6); +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_6); -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__14___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__15(x_1, x_2, x_3, x_4, x_5, x_14, x_15, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_16; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_1); +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_1); lean_dec(x_1); -x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___lambda__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__3(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_13; -x_13 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_4); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_13; +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1() { _start: { -lean_object* x_12; -x_12 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__16___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_2); -lean_dec(x_1); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2() { _start: { -lean_object* x_15; -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_15; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Attr", 4, 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3() { _start: { -lean_object* x_14; -x_14 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__18___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_14; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEq", 7, 7); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4() { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__20(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5() { _start: { -lean_object* x_11; -x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_11; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindFwd", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6() { _start: { -lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_14; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7() { _start: { -lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grindEqRhs", 10, 10); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8() { _start: { -lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(" : ", 3, 3); +x_1 = lean_mk_string_unchecked("grindEqBoth", 11, 11); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2; +x_4 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind(lean_object* x_1) { _start: { -lean_object* x_11; -lean_inc(x_1); -x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Lean_Syntax_getArg(x_1, x_2); +x_4 = l_Lean_Syntax_isNone(x_3); +if (x_4 == 0) { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Lean_Syntax_getArg(x_3, x_5); +lean_dec(x_3); +x_7 = l_Lean_Syntax_getKind(x_6); +x_8 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4; +x_9 = lean_name_eq(x_7, x_8); +if (x_9 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_13 = lean_ctor_get(x_11, 0); -x_14 = l_Lean_MessageData_ofExpr(x_1); -lean_inc(x_2); -x_15 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_15, 0, x_2); -lean_ctor_set(x_15, 1, x_14); -x_16 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; -x_17 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_MessageData_ofExpr(x_13); -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_2); -x_21 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_21, 0, x_4); -lean_ctor_set(x_21, 1, x_20); -x_22 = lean_box(x_3); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_11, 0, x_24); -return x_11; +lean_object* x_10; uint8_t x_11; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6; +x_11 = lean_name_eq(x_7, x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8; +x_13 = lean_name_eq(x_7, x_12); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10; +x_15 = lean_name_eq(x_7, x_14); +lean_dec(x_7); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = 4; +return x_16; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_25 = lean_ctor_get(x_11, 0); -x_26 = lean_ctor_get(x_11, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_11); -x_27 = l_Lean_MessageData_ofExpr(x_1); -lean_inc(x_2); -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_2); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___closed__2; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean_MessageData_ofExpr(x_25); -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_2); -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_4); -lean_ctor_set(x_34, 1, x_33); -x_35 = lean_box(x_3); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_26); -return x_38; +uint8_t x_17; +x_17 = 2; +return x_17; } } else { -uint8_t x_39; -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_39 = !lean_is_exclusive(x_11); -if (x_39 == 0) -{ -return x_11; +uint8_t x_18; +lean_dec(x_7); +x_18 = 1; +return x_18; +} +} +else +{ +uint8_t x_19; +lean_dec(x_7); +x_19 = 3; +return x_19; +} +} +else +{ +uint8_t x_20; +lean_dec(x_7); +x_20 = 0; +return x_20; +} +} +else +{ +uint8_t x_21; +lean_dec(x_3); +x_21 = 5; +return x_21; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_6, 6); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_st_ref_take(x_7, x_8); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 4); +lean_dec(x_15); +x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); +x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; +lean_ctor_set(x_11, 4, x_17); +lean_ctor_set(x_11, 0, x_16); +x_18 = lean_st_ref_set(x_7, x_11, x_12); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_st_ref_take(x_5, x_19); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_21, 1); +lean_dec(x_24); +x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +lean_ctor_set(x_21, 1, x_25); +x_26 = lean_st_ref_set(x_5, x_21, x_22); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +x_29 = lean_box(0); +lean_ctor_set(x_26, 0, x_29); +return x_26; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_26, 1); +lean_inc(x_30); +lean_dec(x_26); +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +return x_32; +} } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_11, 0); -x_41 = lean_ctor_get(x_11, 1); -lean_inc(x_41); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = lean_ctor_get(x_21, 0); +x_34 = lean_ctor_get(x_21, 2); +x_35 = lean_ctor_get(x_21, 3); +x_36 = lean_ctor_get(x_21, 4); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_21); +x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_34); +lean_ctor_set(x_38, 3, x_35); +lean_ctor_set(x_38, 4, x_36); +x_39 = lean_st_ref_set(x_5, x_38, x_22); +x_40 = lean_ctor_get(x_39, 1); lean_inc(x_40); -lean_dec(x_11); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); } +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; } +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\n", 1, 1); -return x_1; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +x_46 = lean_ctor_get(x_11, 2); +x_47 = lean_ctor_get(x_11, 3); +x_48 = lean_ctor_get(x_11, 5); +x_49 = lean_ctor_get(x_11, 6); +x_50 = lean_ctor_get(x_11, 7); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_11); +x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); +x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1; +x_53 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_45); +lean_ctor_set(x_53, 2, x_46); +lean_ctor_set(x_53, 3, x_47); +lean_ctor_set(x_53, 4, x_52); +lean_ctor_set(x_53, 5, x_48); +lean_ctor_set(x_53, 6, x_49); +lean_ctor_set(x_53, 7, x_50); +x_54 = lean_st_ref_set(x_7, x_53, x_12); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_st_ref_take(x_5, x_55); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = lean_ctor_get(x_57, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_57, 2); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 3); +lean_inc(x_61); +x_62 = lean_ctor_get(x_57, 4); +lean_inc(x_62); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + lean_ctor_release(x_57, 2); + lean_ctor_release(x_57, 3); + lean_ctor_release(x_57, 4); + x_63 = x_57; +} else { + lean_dec_ref(x_57); + x_63 = lean_box(0); } +x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 5, 0); +} else { + x_65 = x_63; } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_60); +lean_ctor_set(x_65, 3, x_61); +lean_ctor_set(x_65, 4, x_62); +x_66 = lean_st_ref_set(x_5, x_65, x_58); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_68 = x_66; +} else { + lean_dec_ref(x_66); + x_68 = lean_box(0); +} +x_69 = lean_box(0); +if (lean_is_scalar(x_68)) { + x_70 = lean_alloc_ctor(0, 2, 0); +} else { + x_70 = x_68; } +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_67); +return x_70; } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__2; -x_2 = l_Lean_MessageData_ofFormat(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_5, 1); -x_16 = lean_nat_dec_lt(x_7, x_15); -if (x_16 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_6, x_5); +if (x_13 == 0) { -lean_object* x_17; -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_14; lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_14); -return x_17; +lean_dec(x_9); +lean_dec(x_8); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_6); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_6, 0); -x_20 = lean_ctor_get(x_6, 1); -x_21 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_5, 2); -x_23 = lean_nat_add(x_7, x_22); +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_dec(x_7); -x_7 = x_23; -x_8 = lean_box(0); -x_9 = lean_box(0); -goto _start; -} -else -{ -lean_object* x_25; uint8_t x_26; -lean_free_object(x_6); -x_25 = lean_array_fget(x_2, x_7); -x_26 = lean_unbox(x_19); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_27 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_20); -lean_ctor_set(x_28, 1, x_27); -x_29 = lean_box(0); -x_30 = lean_unbox(x_19); -lean_dec(x_19); -lean_inc(x_13); -lean_inc(x_12); +x_15 = lean_array_uget(x_4, x_6); +x_16 = 1; lean_inc(x_11); lean_inc(x_10); -lean_inc(x_3); -x_31 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_30, x_28, x_29, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -if (lean_obj_tag(x_32) == 0) +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_15, x_16, x_16, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +lean_inc(x_10); +x_21 = l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1(x_20, x_18, x_1, x_8, x_9, x_10, x_11, x_19); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = 1; +x_24 = lean_usize_add(x_6, x_23); +x_25 = lean_box(0); +x_6 = x_24; +x_7 = x_25; +x_12 = x_22; +goto _start; +} +else { -uint8_t x_33; -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_27; lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_33 = !lean_is_exclusive(x_31); -if (x_33 == 0) +lean_dec(x_9); +lean_dec(x_8); +x_27 = !lean_is_exclusive(x_17); +if (x_27 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_31, 0); -lean_dec(x_34); -x_35 = lean_ctor_get(x_32, 0); -lean_inc(x_35); -lean_dec(x_32); -lean_ctor_set(x_31, 0, x_35); -return x_31; +return x_17; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_31, 1); -lean_inc(x_36); -lean_dec(x_31); -x_37 = lean_ctor_get(x_32, 0); -lean_inc(x_37); -lean_dec(x_32); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_36); -return x_38; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_17, 0); +x_29 = lean_ctor_get(x_17, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_17); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_31, 1); -lean_inc(x_39); -lean_dec(x_31); -x_40 = lean_ctor_get(x_32, 0); -lean_inc(x_40); -lean_dec(x_32); -x_41 = lean_ctor_get(x_5, 2); -x_42 = lean_nat_add(x_7, x_41); -lean_dec(x_7); -x_6 = x_40; -x_7 = x_42; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_39; -goto _start; } } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -uint8_t x_44; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_44 = !lean_is_exclusive(x_31); -if (x_44 == 0) +lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_box(0); +x_10 = lean_array_size(x_1); +x_11 = 0; +x_12 = lean_box(0); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2(x_2, x_1, x_9, x_1, x_10, x_11, x_12, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_13) == 0) { -return x_31; +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_12); +return x_13; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_31, 0); -x_46 = lean_ctor_get(x_31, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_31); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; -} +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_12); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } else { -uint8_t x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_19); -x_48 = 0; -x_49 = lean_box(0); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_3); -x_50 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_25, x_3, x_48, x_20, x_49, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_50) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) { -lean_object* x_51; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -if (lean_obj_tag(x_51) == 0) +return x_13; +} +else { -uint8_t x_52; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 0); +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_inc(x_19); lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_52 = !lean_is_exclusive(x_50); -if (x_52 == 0) +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1() { +_start: { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_50, 0); -lean_dec(x_53); -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -lean_ctor_set(x_50, 0, x_54); -return x_50; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` attribute can only be applied to equational theorems or function definitions", 78, 78); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2() { +_start: { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_50, 1); -lean_inc(x_55); -lean_dec(x_50); -x_56 = lean_ctor_get(x_51, 0); -lean_inc(x_56); -lean_dec(x_51); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_55); -return x_57; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is a definition, you must only use the left-hand side for extracting patterns", 79, 79); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3() { +_start: { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_58 = lean_ctor_get(x_50, 1); -lean_inc(x_58); -lean_dec(x_50); -x_59 = lean_ctor_get(x_51, 0); -lean_inc(x_59); -lean_dec(x_51); -x_60 = lean_ctor_get(x_5, 2); -x_61 = lean_nat_add(x_7, x_60); -lean_dec(x_7); -x_6 = x_59; -x_7 = x_61; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_58; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(lean_object* x_1, uint8_t x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -uint8_t x_63; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_10; +lean_inc(x_1); +x_10 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); lean_dec(x_10); +x_13 = l_Lean_ConstantInfo_isTheorem(x_11); +lean_dec(x_11); +if (x_13 == 0) +{ +lean_object* x_14; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_14 = l_Lean_Meta_getEqnsFor_x3f(x_1, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_1); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(x_3); +x_18 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3; +x_19 = lean_string_append(x_18, x_17); +lean_dec(x_17); +x_20 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1; +x_21 = lean_string_append(x_19, x_20); +x_22 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = l_Lean_MessageData_ofFormat(x_22); +x_24 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_23, x_5, x_6, x_7, x_8, x_16); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_3); -x_63 = !lean_is_exclusive(x_50); -if (x_63 == 0) +lean_dec(x_6); +lean_dec(x_5); +return x_24; +} +else +{ +if (x_4 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +lean_dec(x_15); +x_25 = lean_ctor_get(x_14, 1); +lean_inc(x_25); +lean_dec(x_14); +x_26 = l_Lean_MessageData_ofName(x_1); +x_27 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_30, x_5, x_6, x_7, x_8, x_25); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -return x_50; +return x_31; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_50, 0); -x_65 = lean_ctor_get(x_50, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_50); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_31); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_1); +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); +lean_dec(x_14); +x_37 = lean_ctor_get(x_15, 0); +lean_inc(x_37); +lean_dec(x_15); +x_38 = lean_box(0); +x_39 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1(x_37, x_2, x_38, x_5, x_6, x_7, x_8, x_36); +lean_dec(x_37); +return x_39; } } } else { -lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_67 = lean_ctor_get(x_6, 0); -x_68 = lean_ctor_get(x_6, 1); -lean_inc(x_68); -lean_inc(x_67); +uint8_t x_40; +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_69 = l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(x_7, x_1); -if (x_69 == 0) +lean_dec(x_5); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_14); +if (x_40 == 0) { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_67); -lean_ctor_set(x_70, 1, x_68); -x_71 = lean_ctor_get(x_5, 2); -x_72 = lean_nat_add(x_7, x_71); -lean_dec(x_7); -x_6 = x_70; -x_7 = x_72; -x_8 = lean_box(0); -x_9 = lean_box(0); -goto _start; +return x_14; } else { -lean_object* x_74; uint8_t x_75; -x_74 = lean_array_fget(x_2, x_7); -x_75 = lean_unbox(x_67); -if (x_75 == 0) +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_14, 0); +x_42 = lean_ctor_get(x_14, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_14); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; -x_76 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3; -x_77 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_77, 0, x_68); -lean_ctor_set(x_77, 1, x_76); -x_78 = lean_box(0); -x_79 = lean_unbox(x_67); -lean_dec(x_67); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_3); -x_80 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_79, x_77, x_78, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_80) == 0) +uint8_t x_44; lean_object* x_45; +x_44 = 1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_45 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_1, x_44, x_4, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_81; -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -if (lean_obj_tag(x_81) == 0) +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_49 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_48, x_46, x_2, x_5, x_6, x_7, x_8, x_47); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +return x_49; +} +else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_50; +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_3); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_83 = x_80; -} else { - lean_dec_ref(x_80); - x_83 = lean_box(0); +lean_dec(x_6); +lean_dec(x_5); +x_50 = !lean_is_exclusive(x_45); +if (x_50 == 0) +{ +return x_45; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_45, 0); +x_52 = lean_ctor_get(x_45, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_45); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} } -x_84 = lean_ctor_get(x_81, 0); -lean_inc(x_84); -lean_dec(x_81); -if (lean_is_scalar(x_83)) { - x_85 = lean_alloc_ctor(0, 2, 0); -} else { - x_85 = x_83; } -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_82); -return x_85; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_86 = lean_ctor_get(x_80, 1); -lean_inc(x_86); -lean_dec(x_80); -x_87 = lean_ctor_get(x_81, 0); -lean_inc(x_87); -lean_dec(x_81); -x_88 = lean_ctor_get(x_5, 2); -x_89 = lean_nat_add(x_7, x_88); +uint8_t x_54; +lean_dec(x_8); lean_dec(x_7); -x_6 = x_87; -x_7 = x_89; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_86; -goto _start; -} +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_54 = !lean_is_exclusive(x_10); +if (x_54 == 0) +{ +return x_10; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_10, 0); +x_56 = lean_ctor_get(x_10, 1); +lean_inc(x_56); +lean_inc(x_55); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_91 = lean_ctor_get(x_80, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_80, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_93 = x_80; -} else { - lean_dec_ref(x_80); - x_93 = lean_box(0); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } -if (lean_is_scalar(x_93)) { - x_94 = lean_alloc_ctor(1, 2, 0); -} else { - x_94 = x_93; } -lean_ctor_set(x_94, 0, x_91); -lean_ctor_set(x_94, 1, x_92); -return x_94; } } -else +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -uint8_t x_95; lean_object* x_96; lean_object* x_97; -lean_dec(x_67); -x_95 = 0; -x_96 = lean_box(0); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_3); -x_97 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_74, x_3, x_95, x_68, x_96, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_97) == 0) +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); +lean_dec(x_3); +x_10 = l_Lean_ScopedEnvExtension_add___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__1(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_98; -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -if (lean_obj_tag(x_98) == 0) +uint8_t x_13; size_t x_14; size_t x_15; lean_object* x_16; +x_13 = lean_unbox(x_1); +lean_dec(x_1); +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___spec__2(x_13, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___lambda__1(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_100 = x_97; -} else { - lean_dec_ref(x_97); - x_100 = lean_box(0); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = lean_unbox(x_4); +lean_dec(x_4); +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_10, x_11, x_12, x_5, x_6, x_7, x_8, x_9); +return x_13; } -x_101 = lean_ctor_get(x_98, 0); -lean_inc(x_101); -lean_dec(x_98); -if (lean_is_scalar(x_100)) { - x_102 = lean_alloc_ctor(0, 2, 0); -} else { - x_102 = x_100; } -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_99); -return x_102; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`@", 2, 2); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2() { +_start: { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_103 = lean_ctor_get(x_97, 1); -lean_inc(x_103); -lean_dec(x_97); -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -lean_dec(x_98); -x_105 = lean_ctor_get(x_5, 2); -x_106 = lean_nat_add(x_7, x_105); -lean_dec(x_7); -x_6 = x_104; -x_7 = x_106; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_14 = x_103; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3() { +_start: { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_3); -x_108 = lean_ctor_get(x_97, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_97, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_110 = x_97; -} else { - lean_dec_ref(x_97); - x_110 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" theorem ", 9, 9); +return x_1; } -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(1, 2, 0); -} else { - x_111 = x_110; } -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -return x_111; +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` ", 2, 2); +return x_1; +} } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", consider using different options or the `grind_pattern` command", 65, 65); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 1; -x_2 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_3 = lean_box(x_1); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -return x_4; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_9 = lean_array_get_size(x_2); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_9); -lean_ctor_set(x_12, 2, x_11); -x_13 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_13, x_12, x_12, x_14, x_10, lean_box(0), lean_box(0), x_4, x_5, x_6, x_7, x_8); -lean_dec(x_12); +uint8_t x_9; uint8_t x_10; +x_9 = 0; +x_10 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_3, x_9); +if (x_10 == 0) +{ +uint8_t x_11; uint8_t x_12; +x_11 = 1; +x_12 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_3, x_11); +if (x_12 == 0) +{ +uint8_t x_13; uint8_t x_14; +x_13 = 2; +x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_beqTheoremKind____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_6540_(x_3, x_13); +if (x_14 == 0) +{ +lean_object* x_15; +lean_inc(x_1); +x_15 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_15) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_16; lean_object* x_17; uint8_t x_18; x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); +x_18 = l_Lean_ConstantInfo_isTheorem(x_16); lean_dec(x_16); -x_19 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_18, x_4, x_5, x_6, x_7, x_17); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_19; -} -else -{ -uint8_t x_20; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) +if (x_18 == 0) { -return x_15; +uint8_t x_19; lean_object* x_20; +x_19 = 1; +x_20 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_19, x_4, x_5, x_6, x_7, x_17); +return x_20; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +lean_object* x_21; +lean_inc(x_1); +x_21 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor(x_1, x_6, x_7, x_17); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_9; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_1); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_1); +x_25 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_9 = lean_infer_type(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_26 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f(x_24, x_25, x_22, x_3, x_4, x_5, x_6, x_7, x_23); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_2); -x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed), 8, 1); -lean_closure_set(x_13, 0, x_3); -x_14 = 0; -x_15 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_arrowDomainsN___spec__6___rarg(x_10, x_12, x_13, x_14, x_4, x_5, x_6, x_7, x_11); -return x_15; -} -else +lean_object* x_27; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +if (lean_obj_tag(x_27) == 0) { -uint8_t x_16; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute(x_3); +x_30 = l_Lean_stringToMessageData(x_29); +lean_dec(x_29); +x_31 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_MessageData_ofName(x_1); +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6; +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure(x_3); +x_40 = l_Lean_stringToMessageData(x_39); +lean_dec(x_39); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_40); +x_42 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_43, x_4, x_5, x_6, x_7, x_28); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_16 = !lean_is_exclusive(x_9); -if (x_16 == 0) -{ -return x_9; +return x_44; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_9, 0); -x_18 = lean_ctor_get(x_9, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_9); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_3); -lean_dec(x_3); -x_12 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___lambda__1(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_5); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; -x_15 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_1); +x_45 = lean_ctor_get(x_26, 1); +lean_inc(x_45); +lean_dec(x_26); +x_46 = lean_ctor_get(x_27, 0); +lean_inc(x_46); +lean_dec(x_27); +x_47 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_48 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_47, x_46, x_2, x_4, x_5, x_6, x_7, x_45); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_15; +return x_48; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); +uint8_t x_49; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +x_49 = !lean_is_exclusive(x_26); +if (x_49 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage___spec__1___closed__1; -x_8 = lean_panic_fn(x_7, x_1); -x_9 = lean_apply_5(x_8, x_2, x_3, x_4, x_5, x_6); -return x_9; -} +return x_26; } -static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_26, 0); +x_51 = lean_ctor_get(x_26, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_26); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } -static lean_object* _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2; -x_2 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -lean_ctor_set(x_2, 2, x_1); -lean_ctor_set(x_2, 3, x_1); -lean_ctor_set(x_2, 4, x_1); -lean_ctor_set(x_2, 5, x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = lean_ctor_get(x_6, 6); -lean_inc(x_9); +uint8_t x_53; +lean_dec(x_7); lean_dec(x_6); -x_10 = lean_st_ref_take(x_7, x_8); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_14 = lean_ctor_get(x_11, 0); -x_15 = lean_ctor_get(x_11, 4); -lean_dec(x_15); -x_16 = l_Lean_ScopedEnvExtension_addCore___rarg(x_14, x_1, x_2, x_3, x_9); -x_17 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; -lean_ctor_set(x_11, 4, x_17); -lean_ctor_set(x_11, 0, x_16); -x_18 = lean_st_ref_set(x_7, x_11, x_12); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_st_ref_take(x_5, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = !lean_is_exclusive(x_21); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_ctor_get(x_21, 1); -lean_dec(x_24); -x_25 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; -lean_ctor_set(x_21, 1, x_25); -x_26 = lean_st_ref_set(x_5, x_21, x_22); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_26, 0); -lean_dec(x_28); -x_29 = lean_box(0); -lean_ctor_set(x_26, 0, x_29); -return x_26; -} -else +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_53 = !lean_is_exclusive(x_21); +if (x_53 == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_26, 1); -lean_inc(x_30); -lean_dec(x_26); -x_31 = lean_box(0); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -return x_32; -} +return x_21; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_33 = lean_ctor_get(x_21, 0); -x_34 = lean_ctor_get(x_21, 2); -x_35 = lean_ctor_get(x_21, 3); -x_36 = lean_ctor_get(x_21, 4); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_21, 0); +x_55 = lean_ctor_get(x_21, 1); +lean_inc(x_55); +lean_inc(x_54); lean_dec(x_21); -x_37 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; -x_38 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_38, 0, x_33); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_34); -lean_ctor_set(x_38, 3, x_35); -lean_ctor_set(x_38, 4, x_36); -x_39 = lean_st_ref_set(x_5, x_38, x_22); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; -} else { - lean_dec_ref(x_39); - x_41 = lean_box(0); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } -x_42 = lean_box(0); -if (lean_is_scalar(x_41)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_41; } -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; } } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_44 = lean_ctor_get(x_11, 0); -x_45 = lean_ctor_get(x_11, 1); -x_46 = lean_ctor_get(x_11, 2); -x_47 = lean_ctor_get(x_11, 3); -x_48 = lean_ctor_get(x_11, 5); -x_49 = lean_ctor_get(x_11, 6); -x_50 = lean_ctor_get(x_11, 7); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_11); -x_51 = l_Lean_ScopedEnvExtension_addCore___rarg(x_44, x_1, x_2, x_3, x_9); -x_52 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1; -x_53 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_45); -lean_ctor_set(x_53, 2, x_46); -lean_ctor_set(x_53, 3, x_47); -lean_ctor_set(x_53, 4, x_52); -lean_ctor_set(x_53, 5, x_48); -lean_ctor_set(x_53, 6, x_49); -lean_ctor_set(x_53, 7, x_50); -x_54 = lean_st_ref_set(x_7, x_53, x_12); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = lean_st_ref_take(x_5, x_55); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = lean_ctor_get(x_57, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_57, 2); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 3); -lean_inc(x_61); -x_62 = lean_ctor_get(x_57, 4); -lean_inc(x_62); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - lean_ctor_release(x_57, 2); - lean_ctor_release(x_57, 3); - lean_ctor_release(x_57, 4); - x_63 = x_57; -} else { - lean_dec_ref(x_57); - x_63 = lean_box(0); -} -x_64 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2; -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 5, 0); -} else { - x_65 = x_63; -} -lean_ctor_set(x_65, 0, x_59); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_60); -lean_ctor_set(x_65, 3, x_61); -lean_ctor_set(x_65, 4, x_62); -x_66 = lean_st_ref_set(x_5, x_65, x_58); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; -} else { - lean_dec_ref(x_66); - x_68 = lean_box(0); -} -x_69 = lean_box(0); -if (lean_is_scalar(x_68)) { - x_70 = lean_alloc_ctor(0, 2, 0); -} else { - x_70 = x_68; +uint8_t x_57; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_15); +if (x_57 == 0) +{ +return x_15; } -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_67); -return x_70; +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_15, 0); +x_59 = lean_ctor_get(x_15, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_15); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(lean_object* x_1, lean_object* x_2) { -_start: +else { -if (lean_obj_tag(x_1) == 0) +uint8_t x_61; lean_object* x_62; +x_61 = 1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_62 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_61, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_62) == 0) { -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; +lean_object* x_63; uint8_t x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = 0; +x_65 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_64, x_4, x_5, x_6, x_7, x_63); +return x_65; } else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +uint8_t x_66; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_62); +if (x_66 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = l_Lean_Meta_Grind_ppPattern(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); +return x_62; +} +else { -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_62, 0); +x_68 = lean_ctor_get(x_62, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_62); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} } -goto _start; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = l_Lean_Meta_Grind_ppPattern(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; +uint8_t x_70; lean_object* x_71; +x_70 = 0; +x_71 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_70, x_4, x_5, x_6, x_7, x_8); +return x_71; } } +else +{ +uint8_t x_72; lean_object* x_73; +x_72 = 1; +x_73 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr(x_1, x_2, x_3, x_72, x_4, x_5, x_6, x_7, x_8); +return x_73; +} } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; +uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; } -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -lean_object* x_5; -x_5 = lean_ctor_get(x_1, 1); -lean_ctor_set(x_1, 1, x_2); +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) { -lean_object* _tmp_0 = x_5; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = l_Lean_Meta_Grind_EMatchTheorems_erase(x_4, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_2 = x_10; +x_4 = x_8; goto _start; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_1, 0); -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_1); -x_9 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_9, 0, x_7); -lean_ctor_set(x_9, 1, x_2); -x_1 = x_8; -x_2 = x_9; -goto _start; -} +return x_4; } } } -LEAN_EXPORT uint8_t l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { _start: { -if (lean_obj_tag(x_1) == 2) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) { -uint8_t x_2; -x_2 = 1; -return x_2; +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +lean_inc(x_1); +x_8 = l_Lean_Meta_Grind_EMatchTheorems_contains(x_1, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 1; +return x_9; } else { -uint8_t x_3; -x_3 = 0; -return x_3; +size_t x_10; size_t x_11; +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_3 = x_11; +goto _start; +} +} +else +{ +uint8_t x_13; +lean_dec(x_1); +x_13 = 0; +return x_13; } } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1() { _start: { -lean_object* x_1; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt; -return x_1; +uint8_t x_1; uint8_t x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = 0; +x_2 = 1; +x_3 = 1; +x_4 = 0; +x_5 = 2; +x_6 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_6, 0, x_1); +lean_ctor_set_uint8(x_6, 1, x_1); +lean_ctor_set_uint8(x_6, 2, x_1); +lean_ctor_set_uint8(x_6, 3, x_1); +lean_ctor_set_uint8(x_6, 4, x_1); +lean_ctor_set_uint8(x_6, 5, x_2); +lean_ctor_set_uint8(x_6, 6, x_2); +lean_ctor_set_uint8(x_6, 7, x_1); +lean_ctor_set_uint8(x_6, 8, x_2); +lean_ctor_set_uint8(x_6, 9, x_3); +lean_ctor_set_uint8(x_6, 10, x_4); +lean_ctor_set_uint8(x_6, 11, x_2); +lean_ctor_set_uint8(x_6, 12, x_2); +lean_ctor_set_uint8(x_6, 13, x_2); +lean_ctor_set_uint8(x_6, 14, x_5); +lean_ctor_set_uint8(x_6, 15, x_2); +lean_ctor_set_uint8(x_6, 16, x_2); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static uint64_t _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2() { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; -x_13 = lean_array_mk(x_1); -x_14 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_14, 0, x_2); -x_15 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_3); -lean_ctor_set(x_15, 2, x_4); -lean_ctor_set(x_15, 3, x_5); -lean_ctor_set(x_15, 4, x_6); -lean_ctor_set(x_15, 5, x_14); -x_16 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; -x_17 = 0; -x_18 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2(x_16, x_15, x_17, x_8, x_9, x_10, x_11, x_12); -return x_18; +lean_object* x_1; uint64_t x_2; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("invalid pattern(s) for `", 24, 24); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`", 1, 1); -return x_1; +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\nthe following theorem parameters cannot be instantiated:", 57, 57); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint64_t x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1; +x_4 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2; +x_5 = 0; +x_6 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6; +x_7 = l_Lean_Meta_Grind_NormalizePattern_main___closed__1; +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_1); +lean_ctor_set(x_9, 2, x_6); +lean_ctor_set(x_9, 3, x_7); +lean_ctor_set(x_9, 4, x_2); +lean_ctor_set(x_9, 5, x_8); +lean_ctor_set(x_9, 6, x_2); +lean_ctor_set_uint64(x_9, sizeof(void*)*7, x_4); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 8, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 9, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*7 + 10, x_5); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_3 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_2); +lean_ctor_set(x_3, 4, x_2); +lean_ctor_set(x_3, 5, x_2); +lean_ctor_set(x_3, 6, x_2); +lean_ctor_set(x_3, 7, x_2); +lean_ctor_set(x_3, 8, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_2 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 3, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10() { _start: { -lean_object* x_14; -lean_inc(x_12); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8; +x_3 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_4 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5; +x_5 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9; +x_6 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_1); +lean_ctor_set(x_6, 3, x_4); +lean_ctor_set(x_6, 4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_7 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind(x_2); +x_8 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10; +x_9 = lean_st_mk_ref(x_8, x_6); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7; lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_checkCoverage(x_1, x_2, x_3, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_14) == 0) +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr(x_1, x_3, x_7, x_12, x_10, x_4, x_5, x_11); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 0); +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_box(0); -x_18 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2(x_4, x_5, x_1, x_2, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_16); -lean_dec(x_12); +lean_dec(x_13); +x_16 = lean_st_ref_get(x_10, x_15); lean_dec(x_10); -lean_dec(x_9); -return x_18; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_14); +return x_16; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_7); -x_19 = lean_ctor_get(x_14, 1); +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_ctor_get(x_15, 0); -lean_inc(x_20); -lean_dec(x_15); -lean_inc(x_4); -x_21 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_6, x_4); -x_22 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_21, x_4); -x_23 = l_Lean_MessageData_ofList(x_22); -x_24 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_27 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt(x_1, x_2, x_20, x_9, x_10, x_11, x_12, x_19); -if (lean_obj_tag(x_27) == 0) +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_14); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_MessageData_ofName(x_5); -x_31 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4; -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_indentD(x_26); -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6; -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_indentD(x_28); -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_24); -x_42 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_41, x_9, x_10, x_11, x_12, x_29); -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_21; lean_dec(x_10); -lean_dec(x_9); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +x_21 = !lean_is_exclusive(x_13); +if (x_21 == 0) { -return x_42; +return x_13; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_42); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_13); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_array_get_size(x_1); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_lt(x_4, x_3); +if (x_5 == 0) +{ +lean_dec(x_3); +return x_2; } +else +{ +uint8_t x_6; +x_6 = lean_nat_dec_le(x_3, x_3); +if (x_6 == 0) +{ +lean_dec(x_3); +return x_2; } else { -uint8_t x_47; +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = 0; +x_8 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1(x_1, x_7, x_8, x_2); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_take(x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_12, 0); +x_16 = lean_ctor_get(x_12, 4); +lean_dec(x_16); +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed), 2, 1); +lean_closure_set(x_17, 0, x_1); +x_18 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_19 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_18, x_15, x_17); +lean_inc(x_2); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_12, 4, x_10); +lean_ctor_set(x_12, 0, x_19); +x_20 = lean_st_ref_set(x_8, x_12, x_14); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_st_ref_take(x_6, x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_23, 1); lean_dec(x_26); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_5); -x_47 = !lean_is_exclusive(x_27); -if (x_47 == 0) +lean_ctor_set(x_23, 1, x_3); +x_27 = lean_st_ref_set(x_6, x_23, x_24); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); return x_27; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_27, 0); -x_49 = lean_ctor_get(x_27, 1); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); lean_dec(x_27); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_34 = lean_ctor_get(x_23, 0); +x_35 = lean_ctor_get(x_23, 2); +x_36 = lean_ctor_get(x_23, 3); +x_37 = lean_ctor_get(x_23, 4); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_23); +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_3); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 4, x_37); +x_39 = lean_st_ref_set(x_6, x_38, x_24); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); } +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; } +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } else { -uint8_t x_51; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_44 = lean_ctor_get(x_10, 1); +x_45 = lean_ctor_get(x_12, 0); +x_46 = lean_ctor_get(x_12, 1); +x_47 = lean_ctor_get(x_12, 2); +x_48 = lean_ctor_get(x_12, 3); +x_49 = lean_ctor_get(x_12, 5); +x_50 = lean_ctor_get(x_12, 6); +x_51 = lean_ctor_get(x_12, 7); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_51 = !lean_is_exclusive(x_14); -if (x_51 == 0) -{ -return x_14; +x_52 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed), 2, 1); +lean_closure_set(x_52, 0, x_1); +x_53 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_54 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_53, x_45, x_52); +lean_inc(x_2); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 0, x_2); +x_55 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_46); +lean_ctor_set(x_55, 2, x_47); +lean_ctor_set(x_55, 3, x_48); +lean_ctor_set(x_55, 4, x_10); +lean_ctor_set(x_55, 5, x_49); +lean_ctor_set(x_55, 6, x_50); +lean_ctor_set(x_55, 7, x_51); +x_56 = lean_st_ref_set(x_8, x_55, x_44); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_st_ref_take(x_6, x_57); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_ctor_get(x_59, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_59, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_59, 3); +lean_inc(x_63); +x_64 = lean_ctor_get(x_59, 4); +lean_inc(x_64); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + lean_ctor_release(x_59, 2); + lean_ctor_release(x_59, 3); + lean_ctor_release(x_59, 4); + x_65 = x_59; +} else { + lean_dec_ref(x_59); + x_65 = lean_box(0); +} +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(0, 5, 0); +} else { + x_66 = x_65; +} +lean_ctor_set(x_66, 0, x_61); +lean_ctor_set(x_66, 1, x_3); +lean_ctor_set(x_66, 2, x_62); +lean_ctor_set(x_66, 3, x_63); +lean_ctor_set(x_66, 4, x_64); +x_67 = lean_st_ref_set(x_6, x_66, x_60); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_69 = x_67; +} else { + lean_dec_ref(x_67); + x_69 = lean_box(0); +} +x_70 = lean_box(0); +if (lean_is_scalar(x_69)) { + x_71 = lean_alloc_ctor(0, 2, 0); +} else { + x_71 = x_69; +} +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_68); +return x_71; +} } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_14, 0); -x_53 = lean_ctor_get(x_14, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_14); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_72 = lean_ctor_get(x_10, 0); +x_73 = lean_ctor_get(x_10, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_10); +x_74 = lean_ctor_get(x_72, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +x_76 = lean_ctor_get(x_72, 2); +lean_inc(x_76); +x_77 = lean_ctor_get(x_72, 3); +lean_inc(x_77); +x_78 = lean_ctor_get(x_72, 5); +lean_inc(x_78); +x_79 = lean_ctor_get(x_72, 6); +lean_inc(x_79); +x_80 = lean_ctor_get(x_72, 7); +lean_inc(x_80); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + lean_ctor_release(x_72, 2); + lean_ctor_release(x_72, 3); + lean_ctor_release(x_72, 4); + lean_ctor_release(x_72, 5); + lean_ctor_release(x_72, 6); + lean_ctor_release(x_72, 7); + x_81 = x_72; +} else { + lean_dec_ref(x_72); + x_81 = lean_box(0); +} +x_82 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed), 2, 1); +lean_closure_set(x_82, 0, x_1); +x_83 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_84 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_83, x_74, x_82); +lean_inc(x_2); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_2); +lean_ctor_set(x_85, 1, x_2); +if (lean_is_scalar(x_81)) { + x_86 = lean_alloc_ctor(0, 8, 0); +} else { + x_86 = x_81; +} +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_75); +lean_ctor_set(x_86, 2, x_76); +lean_ctor_set(x_86, 3, x_77); +lean_ctor_set(x_86, 4, x_85); +lean_ctor_set(x_86, 5, x_78); +lean_ctor_set(x_86, 6, x_79); +lean_ctor_set(x_86, 7, x_80); +x_87 = lean_st_ref_set(x_8, x_86, x_73); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = lean_st_ref_take(x_6, x_88); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); +x_92 = lean_ctor_get(x_90, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_90, 2); +lean_inc(x_93); +x_94 = lean_ctor_get(x_90, 3); +lean_inc(x_94); +x_95 = lean_ctor_get(x_90, 4); +lean_inc(x_95); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + lean_ctor_release(x_90, 2); + lean_ctor_release(x_90, 3); + lean_ctor_release(x_90, 4); + x_96 = x_90; +} else { + lean_dec_ref(x_90); + x_96 = lean_box(0); +} +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(0, 5, 0); +} else { + x_97 = x_96; +} +lean_ctor_set(x_97, 0, x_92); +lean_ctor_set(x_97, 1, x_3); +lean_ctor_set(x_97, 2, x_93); +lean_ctor_set(x_97, 3, x_94); +lean_ctor_set(x_97, 4, x_95); +x_98 = lean_st_ref_set(x_6, x_97, x_91); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; +} else { + lean_dec_ref(x_98); + x_100 = lean_box(0); +} +x_101 = lean_box(0); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; } +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` is not a theorem, you cannot assign patterns to non-theorems for the `grind` tactic", 85, 85); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__2() { -_start: +lean_object* x_3; lean_object* x_4; +x_3 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_3, 0, x_1); +x_4 = l_Lean_Meta_Grind_EMatchTheorems_erase(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_take(x_8, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_10, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_12, 0); +x_16 = lean_ctor_get(x_12, 4); +lean_dec(x_16); +x_17 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4), 2, 1); +lean_closure_set(x_17, 0, x_1); +x_18 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_19 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_18, x_15, x_17); +lean_inc(x_2); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_12, 4, x_10); +lean_ctor_set(x_12, 0, x_19); +x_20 = lean_st_ref_set(x_8, x_12, x_14); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_st_ref_take(x_6, x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_23, 1); +lean_dec(x_26); +lean_ctor_set(x_23, 1, x_3); +x_27 = lean_st_ref_set(x_6, x_23, x_24); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); +return x_27; } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed), 1, 0); -return x_1; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__4() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("symbols.all fun s => s matches .const _\n ", 42, 42); -return x_1; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_34 = lean_ctor_get(x_23, 0); +x_35 = lean_ctor_get(x_23, 2); +x_36 = lean_ctor_get(x_23, 3); +x_37 = lean_ctor_get(x_23, 4); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_23); +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_3); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 4, x_37); +x_39 = lean_st_ref_set(x_6, x_38, x_24); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); } +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___lambda__1___closed__3; -x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__4; -x_3 = lean_string_append(x_1, x_2); -return x_3; +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.addEMatchTheorem", 32, 32); -return x_1; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_44 = lean_ctor_get(x_10, 1); +x_45 = lean_ctor_get(x_12, 0); +x_46 = lean_ctor_get(x_12, 1); +x_47 = lean_ctor_get(x_12, 2); +x_48 = lean_ctor_get(x_12, 3); +x_49 = lean_ctor_get(x_12, 5); +x_50 = lean_ctor_get(x_12, 6); +x_51 = lean_ctor_get(x_12, 7); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_12); +x_52 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4), 2, 1); +lean_closure_set(x_52, 0, x_1); +x_53 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_54 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_53, x_45, x_52); +lean_inc(x_2); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 0, x_2); +x_55 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_46); +lean_ctor_set(x_55, 2, x_47); +lean_ctor_set(x_55, 3, x_48); +lean_ctor_set(x_55, 4, x_10); +lean_ctor_set(x_55, 5, x_49); +lean_ctor_set(x_55, 6, x_50); +lean_ctor_set(x_55, 7, x_51); +x_56 = lean_st_ref_set(x_8, x_55, x_44); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_st_ref_take(x_6, x_57); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_ctor_get(x_59, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_59, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_59, 3); +lean_inc(x_63); +x_64 = lean_ctor_get(x_59, 4); +lean_inc(x_64); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + lean_ctor_release(x_59, 2); + lean_ctor_release(x_59, 3); + lean_ctor_release(x_59, 4); + x_65 = x_59; +} else { + lean_dec_ref(x_59); + x_65 = lean_box(0); } +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(0, 5, 0); +} else { + x_66 = x_65; +} +lean_ctor_set(x_66, 0, x_61); +lean_ctor_set(x_66, 1, x_3); +lean_ctor_set(x_66, 2, x_62); +lean_ctor_set(x_66, 3, x_63); +lean_ctor_set(x_66, 4, x_64); +x_67 = lean_st_ref_set(x_6, x_66, x_60); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_69 = x_67; +} else { + lean_dec_ref(x_67); + x_69 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1; -x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__6; -x_3 = lean_unsigned_to_nat(340u); -x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_Meta_Grind_addEMatchTheorem___closed__5; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +x_70 = lean_box(0); +if (lean_is_scalar(x_69)) { + x_71 = lean_alloc_ctor(0, 2, 0); +} else { + x_71 = x_69; +} +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_68); +return x_71; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__8() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ematch", 6, 6); -return x_1; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_72 = lean_ctor_get(x_10, 0); +x_73 = lean_ctor_get(x_10, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_10); +x_74 = lean_ctor_get(x_72, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +x_76 = lean_ctor_get(x_72, 2); +lean_inc(x_76); +x_77 = lean_ctor_get(x_72, 3); +lean_inc(x_77); +x_78 = lean_ctor_get(x_72, 5); +lean_inc(x_78); +x_79 = lean_ctor_get(x_72, 6); +lean_inc(x_79); +x_80 = lean_ctor_get(x_72, 7); +lean_inc(x_80); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + lean_ctor_release(x_72, 2); + lean_ctor_release(x_72, 3); + lean_ctor_release(x_72, 4); + lean_ctor_release(x_72, 5); + lean_ctor_release(x_72, 6); + lean_ctor_release(x_72, 7); + x_81 = x_72; +} else { + lean_dec_ref(x_72); + x_81 = lean_box(0); +} +x_82 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__4), 2, 1); +lean_closure_set(x_82, 0, x_1); +x_83 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_84 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_83, x_74, x_82); +lean_inc(x_2); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_2); +lean_ctor_set(x_85, 1, x_2); +if (lean_is_scalar(x_81)) { + x_86 = lean_alloc_ctor(0, 8, 0); +} else { + x_86 = x_81; +} +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_75); +lean_ctor_set(x_86, 2, x_76); +lean_ctor_set(x_86, 3, x_77); +lean_ctor_set(x_86, 4, x_85); +lean_ctor_set(x_86, 5, x_78); +lean_ctor_set(x_86, 6, x_79); +lean_ctor_set(x_86, 7, x_80); +x_87 = lean_st_ref_set(x_8, x_86, x_73); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = lean_st_ref_take(x_6, x_88); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); +x_92 = lean_ctor_get(x_90, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_90, 2); +lean_inc(x_93); +x_94 = lean_ctor_get(x_90, 3); +lean_inc(x_94); +x_95 = lean_ctor_get(x_90, 4); +lean_inc(x_95); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + lean_ctor_release(x_90, 2); + lean_ctor_release(x_90, 3); + lean_ctor_release(x_90, 4); + x_96 = x_90; +} else { + lean_dec_ref(x_90); + x_96 = lean_box(0); } +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(0, 5, 0); +} else { + x_97 = x_96; +} +lean_ctor_set(x_97, 0, x_92); +lean_ctor_set(x_97, 1, x_3); +lean_ctor_set(x_97, 2, x_93); +lean_ctor_set(x_97, 3, x_94); +lean_ctor_set(x_97, 4, x_95); +x_98 = lean_st_ref_set(x_6, x_97, x_91); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; +} else { + lean_dec_ref(x_98); + x_100 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("pattern", 7, 7); -return x_1; +x_101 = lean_box(0); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; } +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; -x_2 = l_Lean_Meta_Grind_addEMatchTheorem___closed__8; -x_3 = l_Lean_Meta_Grind_addEMatchTheorem___closed__9; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__11() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); +x_1 = lean_mk_string_unchecked("` is not marked with the `[grind]` attribute", 44, 44); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__12() { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addEMatchTheorem___closed__11; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_9; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_27; lean_object* x_28; lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_5 = l_Lean_MessageData_ofName(x_1); +x_6 = l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4; +x_7 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2; +x_9 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +x_10 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10; +x_11 = lean_st_mk_ref(x_10, x_4); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_27 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7; +lean_inc(x_1); +x_28 = l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(x_1, x_27, x_12, x_2, x_3, x_13); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_10; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 2) +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_ConstantInfo_isTheorem(x_29); +lean_dec(x_29); +if (x_31 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_ctor_get(x_10, 0); +lean_object* x_32; +lean_inc(x_3); +lean_inc(x_2); lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); +x_32 = l_Lean_Meta_getEqnsFor_x3f(x_1, x_27, x_12, x_2, x_3, x_30); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_9, x_27, x_12, x_2, x_3, x_34); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_box(0); -x_16 = l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(x_14, x_15); -lean_inc(x_1); -x_17 = l_Lean_Expr_const___override(x_1, x_16); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_18 = l_Lean_Meta_Grind_NormalizePattern_main(x_3, x_4, x_5, x_6, x_7, x_11); -if (lean_obj_tag(x_18) == 0) +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = !lean_is_exclusive(x_19); -if (x_22 == 0) +return x_35; +} +else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_19, 0); -x_24 = lean_ctor_get(x_19, 1); -lean_dec(x_24); -x_25 = !lean_is_exclusive(x_20); -if (x_25 == 0) +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_35, 0); +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_35); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_ctor_get(x_20, 0); -x_27 = lean_ctor_get(x_20, 1); -x_28 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; -lean_inc(x_26); -x_29 = l_List_all___rarg(x_26, x_28); -if (x_29 == 0) +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_40 = lean_ctor_get(x_32, 1); +lean_inc(x_40); +lean_dec(x_32); +x_41 = lean_ctor_get(x_33, 0); +lean_inc(x_41); +lean_dec(x_33); +x_42 = lean_st_ref_get(x_3, x_40); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_47 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_48 = l_Lean_ScopedEnvExtension_getState___rarg(x_46, x_47, x_45); +lean_dec(x_45); +x_49 = lean_array_get_size(x_41); +x_50 = lean_unsigned_to_nat(0u); +x_51 = lean_nat_dec_lt(x_50, x_49); +if (x_51 == 0) { -lean_object* x_30; lean_object* x_31; -lean_free_object(x_20); -lean_dec(x_27); -lean_dec(x_26); -lean_free_object(x_19); -lean_dec(x_23); -lean_dec(x_17); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_49); +lean_dec(x_48); +lean_dec(x_9); +x_52 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_53 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_54 = lean_box(0); +x_55 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3(x_41, x_52, x_53, x_54, x_27, x_12, x_2, x_3, x_44); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_30 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; -x_31 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_30, x_4, x_5, x_6, x_7, x_21); -return x_31; +x_14 = x_55; +goto block_26; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; -x_33 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_32, x_4, x_5, x_6, x_7, x_21); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_unbox(x_34); -lean_dec(x_34); -if (x_35 == 0) +size_t x_56; size_t x_57; uint8_t x_58; +x_56 = 0; +x_57 = lean_usize_of_nat(x_49); +lean_dec(x_49); +x_58 = l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2(x_48, x_41, x_56, x_57); +if (x_58 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -lean_free_object(x_20); -lean_free_object(x_19); -x_36 = lean_ctor_get(x_33, 1); -lean_inc(x_36); -lean_dec(x_33); -x_37 = lean_box(0); -x_38 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_37, x_4, x_5, x_6, x_7, x_36); -return x_38; +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_9); +x_59 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_60 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_61 = lean_box(0); +x_62 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3(x_41, x_59, x_60, x_61, x_27, x_12, x_2, x_3, x_44); +lean_dec(x_3); +lean_dec(x_2); +x_14 = x_62; +goto block_26; } else { -uint8_t x_39; -x_39 = !lean_is_exclusive(x_33); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_40 = lean_ctor_get(x_33, 1); -x_41 = lean_ctor_get(x_33, 0); +lean_object* x_63; uint8_t x_64; lean_dec(x_41); -lean_inc(x_1); -x_42 = l_Lean_MessageData_ofName(x_1); -x_43 = l_Lean_Meta_Grind_ppPattern___closed__4; -lean_ctor_set_tag(x_33, 7); -lean_ctor_set(x_33, 1, x_42); -lean_ctor_set(x_33, 0, x_43); -x_44 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; -lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_44); -lean_ctor_set(x_20, 0, x_33); -lean_inc(x_23); -x_45 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); -x_46 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_45, x_15); -x_47 = l_Lean_MessageData_ofList(x_46); -lean_ctor_set_tag(x_19, 7); -lean_ctor_set(x_19, 1, x_47); -lean_ctor_set(x_19, 0, x_20); -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_19); -lean_ctor_set(x_48, 1, x_43); -x_49 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_32, x_48, x_4, x_5, x_6, x_7, x_40); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_50, x_4, x_5, x_6, x_7, x_51); -lean_dec(x_50); -return x_52; +x_63 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_9, x_27, x_12, x_2, x_3, x_44); +lean_dec(x_3); +lean_dec(x_2); +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) +{ +x_14 = x_63; +goto block_26; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_53 = lean_ctor_get(x_33, 1); -lean_inc(x_53); -lean_dec(x_33); -lean_inc(x_1); -x_54 = l_Lean_MessageData_ofName(x_1); -x_55 = l_Lean_Meta_Grind_ppPattern___closed__4; -x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; -lean_ctor_set_tag(x_20, 7); -lean_ctor_set(x_20, 1, x_57); -lean_ctor_set(x_20, 0, x_56); -lean_inc(x_23); -x_58 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); -x_59 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_58, x_15); -x_60 = l_Lean_MessageData_ofList(x_59); -lean_ctor_set_tag(x_19, 7); -lean_ctor_set(x_19, 1, x_60); -lean_ctor_set(x_19, 0, x_20); -x_61 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_61, 0, x_19); -lean_ctor_set(x_61, 1, x_55); -x_62 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_32, x_61, x_4, x_5, x_6, x_7, x_53); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_27, x_15, x_1, x_23, x_26, x_63, x_4, x_5, x_6, x_7, x_64); +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_63, 0); +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_inc(x_65); lean_dec(x_63); -return x_65; +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +x_14 = x_67; +goto block_26; +} } } } } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_66 = lean_ctor_get(x_20, 0); -x_67 = lean_ctor_get(x_20, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_20); -x_68 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; -lean_inc(x_66); -x_69 = l_List_all___rarg(x_66, x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_67); -lean_dec(x_66); -lean_free_object(x_19); -lean_dec(x_23); -lean_dec(x_17); +uint8_t x_68; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_70 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; -x_71 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_70, x_4, x_5, x_6, x_7, x_21); +x_68 = !lean_is_exclusive(x_32); +if (x_68 == 0) +{ +return x_32; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_32, 0); +x_70 = lean_ctor_get(x_32, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_32); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); return x_71; } +} +} else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; -x_72 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; -x_73 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_72, x_4, x_5, x_6, x_7, x_21); -x_74 = lean_ctor_get(x_73, 0); +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_72 = lean_st_ref_get(x_3, x_30); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); lean_inc(x_74); -x_75 = lean_unbox(x_74); -lean_dec(x_74); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_free_object(x_19); -x_76 = lean_ctor_get(x_73, 1); -lean_inc(x_76); +lean_dec(x_72); +x_75 = lean_ctor_get(x_73, 0); +lean_inc(x_75); lean_dec(x_73); -x_77 = lean_box(0); -x_78 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_67, x_15, x_1, x_23, x_66, x_77, x_4, x_5, x_6, x_7, x_76); -return x_78; +x_76 = l_Lean_Meta_Grind_instInhabitedEMatchTheorems; +x_77 = l_Lean_Meta_Grind_addEMatchTheorem___closed__1; +x_78 = l_Lean_ScopedEnvExtension_getState___rarg(x_76, x_77, x_75); +lean_dec(x_75); +lean_inc(x_1); +x_79 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_79, 0, x_1); +x_80 = l_Lean_Meta_Grind_EMatchTheorems_contains(x_78, x_79); +lean_dec(x_79); +if (x_80 == 0) +{ +lean_object* x_81; uint8_t x_82; +lean_dec(x_1); +x_81 = l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(x_9, x_27, x_12, x_2, x_3, x_74); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_12); +x_82 = !lean_is_exclusive(x_81); +if (x_82 == 0) +{ +return x_81; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_79 = lean_ctor_get(x_73, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_80 = x_73; -} else { - lean_dec_ref(x_73); - x_80 = lean_box(0); -} -lean_inc(x_1); -x_81 = l_Lean_MessageData_ofName(x_1); -x_82 = l_Lean_Meta_Grind_ppPattern___closed__4; -if (lean_is_scalar(x_80)) { - x_83 = lean_alloc_ctor(7, 2, 0); -} else { - x_83 = x_80; - lean_ctor_set_tag(x_83, 7); -} -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_81); -x_84 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; -x_85 = lean_alloc_ctor(7, 2, 0); +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_81, 0); +x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_81); +x_85 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); -lean_inc(x_23); -x_86 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_23, x_15); -x_87 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_86, x_15); -x_88 = l_Lean_MessageData_ofList(x_87); -lean_ctor_set_tag(x_19, 7); -lean_ctor_set(x_19, 1, x_88); -lean_ctor_set(x_19, 0, x_85); -x_89 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_89, 0, x_19); -lean_ctor_set(x_89, 1, x_82); -x_90 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_72, x_89, x_4, x_5, x_6, x_7, x_79); -x_91 = lean_ctor_get(x_90, 0); +return x_85; +} +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +lean_dec(x_9); +x_86 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2; +x_87 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2; +x_88 = lean_box(0); +x_89 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5(x_1, x_86, x_87, x_88, x_27, x_12, x_2, x_3, x_74); +lean_dec(x_3); +lean_dec(x_2); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); lean_inc(x_91); -x_92 = lean_ctor_get(x_90, 1); -lean_inc(x_92); -lean_dec(x_90); -x_93 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_67, x_15, x_1, x_23, x_66, x_91, x_4, x_5, x_6, x_7, x_92); -lean_dec(x_91); -return x_93; +lean_dec(x_89); +x_92 = lean_st_ref_get(x_12, x_91); +lean_dec(x_12); +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) +{ +lean_object* x_94; +x_94 = lean_ctor_get(x_92, 0); +lean_dec(x_94); +lean_ctor_set(x_92, 0, x_90); +return x_92; +} +else +{ +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_92, 1); +lean_inc(x_95); +lean_dec(x_92); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_90); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_94 = lean_ctor_get(x_19, 0); -lean_inc(x_94); +uint8_t x_97; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_97 = !lean_is_exclusive(x_28); +if (x_97 == 0) +{ +return x_28; +} +else +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_28, 0); +x_99 = lean_ctor_get(x_28, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_28); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; +} +} +block_26: +{ +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_st_ref_get(x_12, x_16); +lean_dec(x_12); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); lean_dec(x_19); -x_95 = lean_ctor_get(x_20, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_20, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_20)) { - lean_ctor_release(x_20, 0); - lean_ctor_release(x_20, 1); - x_97 = x_20; -} else { - lean_dec_ref(x_20); - x_97 = lean_box(0); +lean_ctor_set(x_17, 0, x_15); +return x_17; } -x_98 = l_Lean_Meta_Grind_addEMatchTheorem___closed__3; -lean_inc(x_95); -x_99 = l_List_all___rarg(x_95, x_98); -if (x_99 == 0) +else { -lean_object* x_100; lean_object* x_101; -lean_dec(x_97); -lean_dec(x_96); -lean_dec(x_95); -lean_dec(x_94); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); lean_dec(x_17); -lean_dec(x_2); -lean_dec(x_1); -x_100 = l_Lean_Meta_Grind_addEMatchTheorem___closed__7; -x_101 = l_panic___at_Lean_Meta_Grind_addEMatchTheorem___spec__1(x_100, x_4, x_5, x_6, x_7, x_21); -return x_101; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_15); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_12); +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) +{ +return x_14; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_14); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; -x_102 = l_Lean_Meta_Grind_addEMatchTheorem___closed__10; -x_103 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_102, x_4, x_5, x_6, x_7, x_21); -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -x_105 = lean_unbox(x_104); -lean_dec(x_104); -if (x_105 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3() { +_start: { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -lean_dec(x_97); -x_106 = lean_ctor_get(x_103, 1); -lean_inc(x_106); -lean_dec(x_103); -x_107 = lean_box(0); -x_108 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_96, x_15, x_1, x_94, x_95, x_107, x_4, x_5, x_6, x_7, x_106); -return x_108; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4() { +_start: { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_109 = lean_ctor_get(x_103, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_110 = x_103; -} else { - lean_dec_ref(x_103); - x_110 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("initFn", 6, 6); +return x_1; } -lean_inc(x_1); -x_111 = l_Lean_MessageData_ofName(x_1); -x_112 = l_Lean_Meta_Grind_ppPattern___closed__4; -if (lean_is_scalar(x_110)) { - x_113 = lean_alloc_ctor(7, 2, 0); -} else { - x_113 = x_110; - lean_ctor_set_tag(x_113, 7); } -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = l_Lean_Meta_Grind_addEMatchTheorem___closed__12; -if (lean_is_scalar(x_97)) { - x_115 = lean_alloc_ctor(7, 2, 0); -} else { - x_115 = x_97; - lean_ctor_set_tag(x_115, 7); +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -lean_inc(x_94); -x_116 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_94, x_15); -x_117 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_116, x_15); -x_118 = l_Lean_MessageData_ofList(x_117); -x_119 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_119, 0, x_115); -lean_ctor_set(x_119, 1, x_118); -x_120 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_112); -x_121 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_102, x_120, x_4, x_5, x_6, x_7, x_109); -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -x_123 = lean_ctor_get(x_121, 1); -lean_inc(x_123); -lean_dec(x_121); -x_124 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_17, x_2, x_96, x_15, x_1, x_94, x_95, x_122, x_4, x_5, x_6, x_7, x_123); -lean_dec(x_122); -return x_124; } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_@", 2, 2); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8() { +_start: { -uint8_t x_125; -lean_dec(x_17); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_125 = !lean_is_exclusive(x_18); -if (x_125 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9() { +_start: { -return x_18; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10() { +_start: { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_18, 0); -x_127 = lean_ctor_get(x_18, 1); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_18); -x_128 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -return x_128; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10; +x_2 = l_Lean_Meta_Grind_mkOffsetPattern___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12() { +_start: { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -x_129 = lean_ctor_get(x_9, 1); -lean_inc(x_129); -lean_dec(x_9); -x_130 = l_Lean_MessageData_ofName(x_1); -x_131 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4; -x_132 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_130); -x_133 = l_Lean_Meta_Grind_addEMatchTheorem___closed__2; -x_134 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_134, 0, x_132); -lean_ctor_set(x_134, 1, x_133); -x_135 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_134, x_4, x_5, x_6, x_7, x_129); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_135; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13() { +_start: { -uint8_t x_136; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_136 = !lean_is_exclusive(x_9); -if (x_136 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_hyg", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14() { +_start: { -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15() { +_start: { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_9, 0); -x_138 = lean_ctor_get(x_9, 1); -lean_inc(x_138); -lean_inc(x_137); -lean_dec(x_9); -x_139 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_139, 0, x_137); -lean_ctor_set(x_139, 1, x_138); -return x_139; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14; +x_2 = lean_unsigned_to_nat(10279u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_mkGroundPattern___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17() { _start: { -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_3); -lean_dec(x_3); -x_10 = l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("The `[grind]` attribute is used to annotate declarations.When applied to an equational theorem, `[grind =]`, `[grind =_]`, or `[grind _=_]`will mark the theorem for use in heuristic instantiations by the `grind` tactic,\n using respectively the left-hand side, the right-hand side, or both sides of the theorem.When applied to a function, `[grind =]` automatically annotates the equational theorems associated with that function.When applied to a theorem `[grind ←]` will instantiate the theorem whenever it encounters the conclusion of the theorem\n (that is, it will use the theorem for backwards reasoning).When applied to a theorem `[grind →]` will instantiate the theorem whenever it encounters sufficiently many of the propositional hypotheses\n (that is, it will use the theorem for forwards reasoning).The attribute `[grind]` by itself will effectively try `[grind ←]` (if the conclusion is sufficient for instantiation) and then `[grind →]`.The `grind` tactic utilizes annotated theorems to add instances of matching patterns into the local context during proof search.For example, if a theorem `@[grind =] theorem foo_idempotent : foo (foo x) = foo x` is annotated,`grind` will add an instance of this theorem to the local context whenever it encounters the pattern `foo (foo x)`.", 1310, 1302); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__1___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18() { _start: { -uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__1(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17; +x_4 = 1; +x_5 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19() { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -return x_13; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___boxed), 6, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20() { _start: { -lean_object* x_14; -x_14 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); -return x_14; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6), 4, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21() { _start: { -lean_object* x_3; uint8_t x_4; -x_3 = lean_st_ref_get(x_1, x_2); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279_(lean_object* x_1) { +_start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; -x_8 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; -x_9 = l_Lean_ScopedEnvExtension_getState___rarg(x_7, x_8, x_6); -lean_dec(x_6); -lean_ctor_set(x_3, 0, x_9); +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21; +x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } -else +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_ctor_get(x_3, 0); -x_11 = lean_ctor_get(x_3, 1); -lean_inc(x_11); -lean_inc(x_10); +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1; -x_14 = l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1; -x_15 = l_Lean_ScopedEnvExtension_getState___rarg(x_13, x_14, x_12); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_11); -return x_16; +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__1(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} } +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____spec__2(x_1, x_2, x_5, x_6); +lean_dec(x_2); +x_8 = lean_box(x_7); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed), 2, 0); -return x_2; +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_2); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Meta_Grind_getEMatchTheorems___rarg(x_1, x_2); +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__2(x_1, x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEMatchTheorems___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_getEMatchTheorems(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_10; } } +lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Grind_Tactics(uint8_t builtin, lean_object*); lean_object* initialize_Lean_HeadIndex(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_FoldConsts(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_CollectFVars(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_InferType(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Eqns(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_EMatchTheorem(uint8_t builtin, lean_object* w) { lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; +res = initialize_Init_Grind_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Grind_Tactics(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_HeadIndex(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -18175,58 +29800,80 @@ lean_dec_ref(res); res = initialize_Lean_Meta_InferType(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Eqns(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Util(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_mkOffsetPattern___closed__1 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__1); +l_Lean_Meta_Grind_mkOffsetPattern___closed__2 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__2); +l_Lean_Meta_Grind_mkOffsetPattern___closed__3 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__3); +l_Lean_Meta_Grind_mkOffsetPattern___closed__4 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__4); +l_Lean_Meta_Grind_mkOffsetPattern___closed__5 = _init_l_Lean_Meta_Grind_mkOffsetPattern___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_mkOffsetPattern___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_detectOffsets___closed__1); +l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1 = _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__1); +l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2 = _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__2); +l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3 = _init_l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_isOffsetPattern_x3f___closed__3); +l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__1); +l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___lambda__1___closed__2); +l_Lean_Meta_Grind_preprocessPattern___closed__1 = _init_l_Lean_Meta_Grind_preprocessPattern___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___closed__1); +l_Lean_Meta_Grind_preprocessPattern___closed__2 = _init_l_Lean_Meta_Grind_preprocessPattern___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___closed__2); +l_Lean_Meta_Grind_preprocessPattern___closed__3 = _init_l_Lean_Meta_Grind_preprocessPattern___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_preprocessPattern___closed__3); l_Lean_Meta_Grind_instInhabitedOrigin___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedOrigin___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedOrigin___closed__1); l_Lean_Meta_Grind_instInhabitedOrigin = _init_l_Lean_Meta_Grind_instInhabitedOrigin(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedOrigin); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__1); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__2); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__3); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__4); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__5); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__6); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__7); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__8); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__9); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__10); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__11); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__12); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__13); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__14); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__15); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__16); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_51____closed__17); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__8); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__9); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__10); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__11); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__12); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__13); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_reprOrigin____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_614____closed__14); l_Lean_Meta_Grind_instReprOrigin___closed__1 = _init_l_Lean_Meta_Grind_instReprOrigin___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instReprOrigin___closed__1); l_Lean_Meta_Grind_instReprOrigin = _init_l_Lean_Meta_Grind_instReprOrigin(); lean_mark_persistent(l_Lean_Meta_Grind_instReprOrigin); -l_Lean_Meta_Grind_Origin_key___closed__1 = _init_l_Lean_Meta_Grind_Origin_key___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_key___closed__1); -l_Lean_Meta_Grind_Origin_key___closed__2 = _init_l_Lean_Meta_Grind_Origin_key___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_key___closed__2); -l_Lean_Meta_Grind_Origin_pp___rarg___closed__1 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__1); -l_Lean_Meta_Grind_Origin_pp___rarg___closed__2 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__2); -l_Lean_Meta_Grind_Origin_pp___rarg___closed__3 = _init_l_Lean_Meta_Grind_Origin_pp___rarg___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_Origin_pp___rarg___closed__3); +l_Lean_Meta_Grind_instBEqOrigin___closed__1 = _init_l_Lean_Meta_Grind_instBEqOrigin___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqOrigin___closed__1); +l_Lean_Meta_Grind_instBEqOrigin = _init_l_Lean_Meta_Grind_instBEqOrigin(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqOrigin); l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__1); l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__2(); @@ -18235,14 +29882,20 @@ l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3 = _init_l_Lean_Meta_Gri lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem___closed__3); l_Lean_Meta_Grind_instInhabitedEMatchTheorem = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorem(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorem); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__1); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1___closed__2); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1); +l_Lean_Meta_Grind_instInhabitedEMatchTheorems = _init_l_Lean_Meta_Grind_instInhabitedEMatchTheorems(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedEMatchTheorems); l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1(); lean_mark_persistent(l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__1); -l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__1___closed__2); -l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1(); -l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2(); -l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__6___closed__1); +l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__1(); +l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__2(); +l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3(); +lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__3___closed__3); l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__1); l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__2(); @@ -18251,53 +29904,47 @@ l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3 = _init_l_Lean_Meta_Grind_EM lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__3); l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4 = _init_l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4(); lean_mark_persistent(l_Lean_Meta_Grind_EMatchTheorems_insert___closed__4); -l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__1); -l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1___closed__2); -l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____spec__1); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__1); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__2); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__3); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__4); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__5); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__6); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__7); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__8); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__9); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__10); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__11); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__12); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__13); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__14); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__15); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__16); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__17); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__18); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__19); -l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20(); -lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687____closed__20); -if (builtin) {res = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_687_(lean_io_mk_world()); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__2); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__3); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__4); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__5); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__6); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__7); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__9); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__10); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__11); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__12); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__13); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__14); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__15); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__16); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__17); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__18); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__19); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666____closed__20); +if (builtin) {res = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_1666_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt = lean_io_result_get_value(res); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ematchTheoremsExt); @@ -18390,8 +30037,8 @@ l_Lean_Meta_Grind_ppPattern___closed__10 = _init_l_Lean_Meta_Grind_ppPattern___c lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__10); l_Lean_Meta_Grind_ppPattern___closed__11 = _init_l_Lean_Meta_Grind_ppPattern___closed__11(); lean_mark_persistent(l_Lean_Meta_Grind_ppPattern___closed__11); -l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_getPatternFunMask___closed__1); +l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1 = _init_l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_NormalizePattern_getPatternSupportMask___closed__1); l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1(); lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__1); l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_NormalizePattern_go___spec__3___closed__2(); @@ -18466,49 +30113,233 @@ l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0 lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___spec__1___closed__3); l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_ppParamsAt___lambda__1___closed__1); -l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1(); -lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__1); -l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2(); -lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__2___closed__2); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__2___closed__1); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__1); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__2); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__3); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__4); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__5); -l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6 = _init_l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___lambda__3___closed__6); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__1); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__2); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__3); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__4); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__5); +l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___lambda__2___closed__6); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__1); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__2); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__3); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__4); +l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5 = _init_l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchTheoremCore___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getProofFor___closed__2); +l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__1); +l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_mkEMatchEqTheoremCore___lambda__3___closed__2); +l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1(); +lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__1); +l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2(); +lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_Meta_Grind_addEMatchTheorem___spec__1___closed__2); l_Lean_Meta_Grind_addEMatchTheorem___closed__1 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__1); -l_Lean_Meta_Grind_addEMatchTheorem___closed__2 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__2); -l_Lean_Meta_Grind_addEMatchTheorem___closed__3 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__3); -l_Lean_Meta_Grind_addEMatchTheorem___closed__4 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__4); -l_Lean_Meta_Grind_addEMatchTheorem___closed__5 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__5); -l_Lean_Meta_Grind_addEMatchTheorem___closed__6 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__6); -l_Lean_Meta_Grind_addEMatchTheorem___closed__7 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__7); -l_Lean_Meta_Grind_addEMatchTheorem___closed__8 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__8); -l_Lean_Meta_Grind_addEMatchTheorem___closed__9 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__9); -l_Lean_Meta_Grind_addEMatchTheorem___closed__10 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__10(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__10); -l_Lean_Meta_Grind_addEMatchTheorem___closed__11 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__11(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__11); -l_Lean_Meta_Grind_addEMatchTheorem___closed__12 = _init_l_Lean_Meta_Grind_addEMatchTheorem___closed__12(); -lean_mark_persistent(l_Lean_Meta_Grind_addEMatchTheorem___closed__12); -return lean_io_result_mk_ok(lean_box(0)); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_noConfusion___rarg___closed__1); +l_Lean_Meta_Grind_instInhabitedTheoremKind = _init_l_Lean_Meta_Grind_instInhabitedTheoremKind(); +l_Lean_Meta_Grind_instBEqTheoremKind___closed__1 = _init_l_Lean_Meta_Grind_instBEqTheoremKind___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqTheoremKind___closed__1); +l_Lean_Meta_Grind_instBEqTheoremKind = _init_l_Lean_Meta_Grind_instBEqTheoremKind(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqTheoremKind); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_toAttribute___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_TheoremKind_explainFailure___closed__7); +l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___spec__2___closed__1(); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addNewPattern___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__6___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collect___lambda__7___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_collectPatterns_x3f___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_mkEMatchTheoremWithKind_x3f___lambda__2___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__8); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__9); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_getKind___closed__10); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindEqAttr___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__1); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__2); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__3); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__4); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__5); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__6); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__7); +l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_addGrindAttr___closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__2(); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__3); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__4); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__5); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__6); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__7); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__9); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__1___closed__10); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____lambda__6___closed__2); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__1); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__2); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__3); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__4); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__5); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__6); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__7); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__8); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__9); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__10); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__11); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__12); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__13); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__14); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__15); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__16); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__17); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__18); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__19); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__20); +l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21 = _init_l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21(); +lean_mark_persistent(l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279____closed__21); +if (builtin) {res = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_EMatchTheorem___hyg_10279_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c index 65165f8e3b54..55b5847ef533 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/ForallProp.c @@ -15,495 +15,1640 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateImpliesDown___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___closed__3; lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6; -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2; -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___closed__4; +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2; lean_object* l_Lean_Meta_Grind_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5; -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3; -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___closed__1; +static lean_object* l_Lean_Meta_Grind_propagateImpliesDown___closed__5; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4; +static lean_object* l_Lean_Meta_Grind_propagateImpliesDown___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateImpliesDown___closed__6; +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9; +lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4; +lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3; +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8; +static lean_object* l_Lean_Meta_Grind_propagateImpliesDown___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10; +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7; +lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8; +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4; +lean_object* l_Lean_indentExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateImpliesDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1; lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_pushEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateImpliesDown___closed__4; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_isEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2; +static lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5; lean_object* lean_expr_instantiate1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7; lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("of_eq_true", 10, 10); +x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("imp_eq_of_eq_true_right", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +x_1 = lean_mk_string_unchecked("imp_eq_of_eq_true_left", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("forall_propagator", 17, 17); +x_1 = lean_mk_string_unchecked("imp_eq_of_eq_false_left", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4; -x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5; -x_3 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8() { +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_16; -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); +lean_object* x_14; lean_inc(x_1); -x_16 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_16) == 0) +x_14 = l_Lean_Meta_Grind_isEqFalse(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3; +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_14, 1); lean_inc(x_17); +lean_dec(x_14); lean_inc(x_1); -x_20 = l_Lean_mkAppB(x_19, x_1, x_17); -x_21 = lean_expr_instantiate1(x_2, x_20); -lean_dec(x_20); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_9); -x_22 = l_Lean_Meta_Grind_simp(x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_18); +x_18 = l_Lean_Meta_Grind_isEqTrue(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +lean_inc(x_2); +x_22 = l_Lean_Meta_Grind_isEqTrue(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_21); if (lean_obj_tag(x_22) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_23; uint8_t x_24; x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -lean_inc(x_1); -x_25 = l_Lean_Expr_lam___override(x_3, x_1, x_2, x_4); -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_inc(x_5); -x_27 = l_Lean_Meta_Grind_getGeneration(x_5, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_24); -if (lean_obj_tag(x_27) == 0) +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_22, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_22, 0, x_27); +return x_22; +} +else { lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); +x_28 = lean_ctor_get(x_22, 1); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_14); -lean_inc(x_13); +lean_dec(x_22); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_26); -x_30 = l_Lean_Meta_Grind_internalize(x_26, x_28, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_29); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -x_32 = l_Lean_Meta_Simp_Result_getProof(x_23, x_11, x_12, x_13, x_14, x_31); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_32 = l_Lean_Meta_Grind_mkEqTrueProof(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_31); if (lean_obj_tag(x_32) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); lean_dec(x_32); -x_35 = l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8; -lean_inc(x_26); -x_36 = l_Lean_mkApp5(x_35, x_1, x_25, x_26, x_17, x_33); -x_37 = 0; -x_38 = l_Lean_Meta_Grind_pushEqCore(x_5, x_26, x_36, x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_34); -lean_dec(x_14); -lean_dec(x_13); +x_35 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5; +x_36 = l_Lean_mkApp3(x_35, x_1, x_2, x_33); +x_37 = l_Lean_Meta_Grind_pushEqTrue(x_3, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_34); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_38; +lean_dec(x_6); +lean_dec(x_5); +return x_37; } else { -uint8_t x_39; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_38; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_39 = !lean_is_exclusive(x_32); -if (x_39 == 0) +x_38 = !lean_is_exclusive(x_32); +if (x_38 == 0) { return x_32; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_32, 0); -x_41 = lean_ctor_get(x_32, 1); -lean_inc(x_41); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_32, 0); +x_40 = lean_ctor_get(x_32, 1); lean_inc(x_40); +lean_inc(x_39); lean_dec(x_32); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} } } } else { -uint8_t x_43; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_23); -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_42; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_43 = !lean_is_exclusive(x_30); -if (x_43 == 0) +x_42 = !lean_is_exclusive(x_22); +if (x_42 == 0) { -return x_30; +return x_22; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_30, 0); -x_45 = lean_ctor_get(x_30, 1); -lean_inc(x_45); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_22, 0); +x_44 = lean_ctor_get(x_22, 1); lean_inc(x_44); -lean_dec(x_30); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_inc(x_43); +lean_dec(x_22); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_47; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_23); -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_18, 1); +lean_inc(x_46); +lean_dec(x_18); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_47 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_46); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8; +lean_inc(x_2); +x_51 = l_Lean_mkApp3(x_50, x_1, x_2, x_48); +x_52 = 0; +x_53 = l_Lean_Meta_Grind_pushEqCore(x_3, x_2, x_51, x_52, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_49); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_27); -if (x_47 == 0) -{ -return x_27; -} -else -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_27, 0); -x_49 = lean_ctor_get(x_27, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_27); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; -} -} +return x_53; } else { -uint8_t x_51; -lean_dec(x_17); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_54; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_22); -if (x_51 == 0) +x_54 = !lean_is_exclusive(x_47); +if (x_54 == 0) { -return x_22; +return x_47; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_22, 0); -x_53 = lean_ctor_get(x_22, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_22); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_47, 0); +x_56 = lean_ctor_get(x_47, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_47); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} } } } else { -uint8_t x_55; -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_58; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = !lean_is_exclusive(x_16); -if (x_55 == 0) +x_58 = !lean_is_exclusive(x_18); +if (x_58 == 0) { -return x_16; +return x_18; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_16, 0); -x_57 = lean_ctor_get(x_16, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_16); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; -} +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_18, 0); +x_60 = lean_ctor_get(x_18, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_18); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_1) == 7) +else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 2); -lean_inc(x_13); -x_14 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_14, 1); +lean_inc(x_62); +lean_dec(x_14); lean_inc(x_12); -x_15 = l_Lean_Meta_Grind_isEqTrue(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_15) == 0) +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_63 = l_Lean_Meta_Grind_mkEqFalseProof(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_62); +if (lean_obj_tag(x_63) == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11; +x_67 = l_Lean_mkApp3(x_66, x_1, x_2, x_64); +x_68 = l_Lean_Meta_Grind_pushEqTrue(x_3, x_67, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_65); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_68; +} +else { -uint8_t x_18; -lean_dec(x_13); +uint8_t x_69; lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) +x_69 = !lean_is_exclusive(x_63); +if (x_69 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_15, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_15, 0, x_20); -return x_15; +return x_63; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); -lean_dec(x_15); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_63, 0); +x_71 = lean_ctor_get(x_63, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_63); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_15, 1); -lean_inc(x_24); -lean_dec(x_15); -x_25 = lean_box(0); -x_26 = l_Lean_Meta_Grind_propagateForallProp___lambda__1(x_12, x_13, x_11, x_14, x_1, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_24); -return x_26; } } else { -uint8_t x_27; -lean_dec(x_13); +uint8_t x_73; lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_27 = !lean_is_exclusive(x_15); -if (x_27 == 0) +x_73 = !lean_is_exclusive(x_14); +if (x_73 == 0) { -return x_15; +return x_14; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_15, 0); -x_29 = lean_ctor_get(x_15, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_15); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_14, 0); +x_75 = lean_ctor_get(x_14, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_14); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_31; lean_object* x_32; -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = l_Lean_Meta_Grind_alreadyInternalized(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +uint8_t x_16; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1(x_2, x_3, x_1, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_14; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("forall_propagator", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +lean_object* x_18; +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +x_18 = l_Lean_Meta_Simp_Result_getProof(x_1, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2; +x_22 = l_Lean_Expr_const___override(x_21, x_2); +lean_inc(x_5); +x_23 = l_Lean_mkApp5(x_22, x_3, x_4, x_5, x_6, x_19); +x_24 = 0; +x_25 = l_Lean_Meta_Grind_pushEqCore(x_7, x_5, x_23, x_24, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +return x_25; +} +else +{ +uint8_t x_26; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_18); +if (x_26 == 0) +{ +return x_18; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_18); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("of_eq_true", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("q': ", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" for", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_17 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +x_21 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3; +lean_inc(x_18); +lean_inc(x_1); +x_22 = l_Lean_mkAppB(x_21, x_1, x_18); +x_23 = lean_expr_instantiate1(x_2, x_22); +lean_dec(x_22); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_10); +x_24 = l_Lean_Meta_Grind_simp(x_23, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_19); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +lean_inc(x_1); +x_27 = l_Lean_Expr_lam___override(x_3, x_1, x_2, x_4); +x_28 = lean_ctor_get(x_25, 0); +lean_inc(x_28); +lean_inc(x_5); +x_29 = l_Lean_Meta_Grind_getGeneration(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_26); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_28); +x_32 = l_Lean_Meta_Grind_internalize(x_28, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +lean_inc(x_6); +x_34 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_33); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_unbox(x_35); +lean_dec(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_6); +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_box(0); +x_39 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(x_25, x_20, x_1, x_27, x_28, x_18, x_5, x_38, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_37); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_39; +} +else +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_34); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_34, 1); +x_42 = lean_ctor_get(x_34, 0); +lean_dec(x_42); +x_43 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_41); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +lean_inc(x_28); +x_45 = l_Lean_MessageData_ofExpr(x_28); +x_46 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5; +lean_ctor_set_tag(x_34, 7); +lean_ctor_set(x_34, 1, x_45); +lean_ctor_set(x_34, 0, x_46); +x_47 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7; +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_34); +lean_ctor_set(x_48, 1, x_47); +lean_inc(x_5); +x_49 = l_Lean_indentExpr(x_5); +x_50 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +x_53 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_6, x_52, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_44); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(x_25, x_20, x_1, x_27, x_28, x_18, x_5, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_55); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_54); +return x_56; +} +else +{ +uint8_t x_57; +lean_free_object(x_34); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_43); +if (x_57 == 0) +{ +return x_43; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_43, 0); +x_59 = lean_ctor_get(x_43, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_43); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_34, 1); +lean_inc(x_61); +lean_dec(x_34); +x_62 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_61); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +lean_inc(x_28); +x_64 = l_Lean_MessageData_ofExpr(x_28); +x_65 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5; +x_66 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_64); +x_67 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7; +x_68 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +lean_inc(x_5); +x_69 = l_Lean_indentExpr(x_5); +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_6, x_72, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_63); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(x_25, x_20, x_1, x_27, x_28, x_18, x_5, x_74, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_75); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_74); +return x_76; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_77 = lean_ctor_get(x_62, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_62, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_79 = x_62; +} else { + lean_dec_ref(x_62); + x_79 = lean_box(0); +} +if (lean_is_scalar(x_79)) { + x_80 = lean_alloc_ctor(1, 2, 0); +} else { + x_80 = x_79; +} +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_78); +return x_80; +} +} +} +} +else +{ +uint8_t x_81; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_81 = !lean_is_exclusive(x_32); +if (x_81 == 0) +{ +return x_32; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_32, 0); +x_83 = lean_ctor_get(x_32, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_32); +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +return x_84; +} +} +} +else +{ +uint8_t x_85; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_85 = !lean_is_exclusive(x_29); +if (x_85 == 0) +{ +return x_29; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_29, 0); +x_87 = lean_ctor_get(x_29, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_29); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; +} +} +} +else +{ +uint8_t x_89; +lean_dec(x_18); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_89 = !lean_is_exclusive(x_24); +if (x_89 == 0) +{ +return x_24; +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_24, 0); +x_91 = lean_ctor_get(x_24, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_24); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; +} +} +} +else +{ +uint8_t x_93; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_93 = !lean_is_exclusive(x_17); +if (x_93 == 0) +{ +return x_17; +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_17, 0); +x_95 = lean_ctor_get(x_17, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_17); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isEqTrue, ", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_inc(x_1); +x_17 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_box(0); +x_22 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2(x_2, x_3, x_4, x_5, x_6, x_1, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +return x_22; +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 1); +x_25 = lean_ctor_get(x_17, 0); +lean_dec(x_25); +x_26 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +lean_inc(x_6); +x_28 = l_Lean_MessageData_ofExpr(x_6); +x_29 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_28); +lean_ctor_set(x_17, 0, x_29); +x_30 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_17); +lean_ctor_set(x_31, 1, x_30); +lean_inc(x_1); +x_32 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_1, x_31, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_27); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2(x_2, x_3, x_4, x_5, x_6, x_1, x_33, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +lean_dec(x_33); +return x_35; +} +else +{ +uint8_t x_36; +lean_free_object(x_17); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_26); +if (x_36 == 0) +{ +return x_26; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_26, 0); +x_38 = lean_ctor_get(x_26, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_26); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_17, 1); +lean_inc(x_40); +lean_dec(x_17); +x_41 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +lean_inc(x_6); +x_43 = l_Lean_MessageData_ofExpr(x_6); +x_44 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +x_46 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_inc(x_1); +x_48 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_1, x_47, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_42); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2(x_2, x_3, x_4, x_5, x_6, x_1, x_49, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_50); +lean_dec(x_49); +return x_51; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_52 = lean_ctor_get(x_41, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_41, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_54 = x_41; +} else { + lean_dec_ref(x_41); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(1, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = l_Lean_Expr_hasLooseBVars(x_1); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_5); +lean_dec(x_4); +x_18 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp(x_2, x_3, x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_18; +} +else +{ +lean_object* x_19; +lean_inc(x_3); +x_19 = l_Lean_Meta_Grind_isEqTrue(x_3, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_19); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_19, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_19, 0, x_24); +return x_19; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 1); +lean_inc(x_25); +lean_dec(x_19); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_19, 1); +lean_inc(x_28); +lean_dec(x_19); +x_29 = lean_box(0); +x_30 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3(x_4, x_3, x_1, x_5, x_6, x_2, x_29, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_28); +return x_30; +} +} +else +{ +uint8_t x_31; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_19); +if (x_31 == 0) +{ +return x_19; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_19, 0); +x_33 = lean_ctor_get(x_19, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_19); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("forallPropagator", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp___closed__2; +x_3 = l_Lean_Meta_Grind_propagateForallPropUp___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_1) == 7) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +x_15 = l_Lean_Meta_Grind_propagateForallPropUp___closed__4; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_box(0); +x_21 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(x_13, x_1, x_12, x_15, x_11, x_14, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +return x_21; +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 1); +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +x_25 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +lean_inc(x_1); +x_27 = l_Lean_MessageData_ofExpr(x_1); +x_28 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +lean_ctor_set_tag(x_16, 7); +lean_ctor_set(x_16, 1, x_27); +lean_ctor_set(x_16, 0, x_28); +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_16); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(x_13, x_1, x_12, x_15, x_11, x_14, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_32); +lean_dec(x_31); +return x_33; +} +else +{ +uint8_t x_34; +lean_free_object(x_16); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -511,23 +1656,440 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_31 = lean_box(0); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_10); -return x_32; +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) +{ +return x_25; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_ctor_get(x_25, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_25); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_16, 1); +lean_inc(x_38); +lean_dec(x_16); +x_39 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +lean_inc(x_1); +x_41 = l_Lean_MessageData_ofExpr(x_1); +x_42 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(x_13, x_1, x_12, x_15, x_11, x_14, x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_47); +lean_dec(x_46); +return x_48; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_49 = lean_ctor_get(x_39, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_39, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_51 = x_39; +} else { + lean_dec_ref(x_39); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} +} +} +else +{ +lean_object* x_53; lean_object* x_54; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_53 = lean_box(0); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_10); +return x_54; +} } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_18; +} } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallProp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -uint8_t x_16; lean_object* x_17; -x_16 = lean_unbox(x_4); +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_4); lean_dec(x_4); -x_17 = l_Lean_Meta_Grind_propagateForallProp___lambda__1(x_1, x_2, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_18 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__2(x_1, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_5); +lean_dec(x_5); +x_18 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__3(x_1, x_2, x_3, x_4, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateForallPropUp___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_6); lean_dec(x_6); -return x_17; +x_18 = l_Lean_Meta_Grind_propagateForallPropUp___lambda__4(x_1, x_2, x_3, x_4, x_5, x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); +return x_18; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_true_of_imp_eq_false", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateImpliesDown___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateImpliesDown___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_false_of_imp_eq_false", 24, 24); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateImpliesDown___closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateImpliesDown___closed__5; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateImpliesDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_isEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +uint8_t x_14; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_11, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +else +{ +if (lean_obj_tag(x_1) == 7) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_dec(x_11); +x_21 = lean_ctor_get(x_1, 1); +lean_inc(x_21); +x_22 = lean_ctor_get(x_1, 2); +lean_inc(x_22); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_23 = l_Lean_Meta_Grind_mkEqFalseProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_Meta_Grind_propagateImpliesDown___closed__3; +lean_inc(x_24); +lean_inc(x_22); +lean_inc(x_21); +x_27 = l_Lean_mkApp3(x_26, x_21, x_22, x_24); +lean_inc(x_21); +x_28 = l_Lean_Meta_Grind_pushEqTrue(x_21, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = l_Lean_Meta_Grind_propagateImpliesDown___closed__6; +lean_inc(x_22); +x_31 = l_Lean_mkApp3(x_30, x_21, x_22, x_24); +x_32 = l_Lean_Meta_Grind_pushEqFalse(x_22, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_32; +} +else +{ +uint8_t x_33; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_33 = !lean_is_exclusive(x_23); +if (x_33 == 0) +{ +return x_23; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_23, 0); +x_35 = lean_ctor_get(x_23, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_23); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_9); +lean_dec(x_8); +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_11); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_11, 0); +lean_dec(x_38); +x_39 = lean_box(0); +lean_ctor_set(x_11, 0, x_39); +return x_11; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_11, 1); +lean_inc(x_40); +lean_dec(x_11); +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +} +} +else +{ +uint8_t x_43; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_11); +if (x_43 == 0) +{ +return x_11; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_11); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} } } lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); @@ -551,22 +2113,74 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__1); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__2); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__3); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__4); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__5); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__6); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__7); -l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateForallProp___lambda__1___closed__8); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__3); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__4); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__5); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__6); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__7); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__8); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__9); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__10); +l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11 = _init_l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp_propagateImpliesUp___lambda__1___closed__11); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__3); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__4); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__5); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__6); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__7); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__8); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__2___closed__9); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___lambda__3___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp___closed__1 = _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___closed__1); +l_Lean_Meta_Grind_propagateForallPropUp___closed__2 = _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___closed__2); +l_Lean_Meta_Grind_propagateForallPropUp___closed__3 = _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___closed__3); +l_Lean_Meta_Grind_propagateForallPropUp___closed__4 = _init_l_Lean_Meta_Grind_propagateForallPropUp___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateForallPropUp___closed__4); +l_Lean_Meta_Grind_propagateImpliesDown___closed__1 = _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateImpliesDown___closed__1); +l_Lean_Meta_Grind_propagateImpliesDown___closed__2 = _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateImpliesDown___closed__2); +l_Lean_Meta_Grind_propagateImpliesDown___closed__3 = _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateImpliesDown___closed__3); +l_Lean_Meta_Grind_propagateImpliesDown___closed__4 = _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateImpliesDown___closed__4); +l_Lean_Meta_Grind_propagateImpliesDown___closed__5 = _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateImpliesDown___closed__5); +l_Lean_Meta_Grind_propagateImpliesDown___closed__6 = _init_l_Lean_Meta_Grind_propagateImpliesDown___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateImpliesDown___closed__6); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c index c0a9bbeba790..09906c01e6cc 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Internalize.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Internalize -// Imports: Init.Grind.Util Lean.Meta.LitValues Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util +// Imports: Init.Grind.Util Init.Grind.Lemmas Lean.Meta.LitValues Lean.Meta.Match.MatcherInfo Lean.Meta.Match.MatchEqsExt Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Util #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,233 +13,310 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +uint8_t lean_is_matcher(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3; +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__5; extern lean_object* l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1; +uint8_t l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__1; -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__9; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2; +uint8_t l_Lean_Expr_isDIte(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3; lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__4; +LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_addCongrTable___spec__6(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___boxed(lean_object**); lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); lean_object* l_Lean_Meta_Grind_mkENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4; +uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__4; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__3; +lean_object* l_Lean_PersistentHashMap_insert___at_Lean_NameSSet_insert___spec__2(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Meta_Grind_congrHash(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes; lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9; lean_object* l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4; +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1; static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__2; lean_object* l_Lean_Meta_Grind_setENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(lean_object*, size_t, lean_object*); +uint8_t l_Lean_Meta_Grind_EMatchTheorems_isErased(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_propagateUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isBVar(lean_object*); extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; lean_object* l_Lean_Meta_Grind_EMatchTheorems_insert(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2; +lean_object* l_Lean_mkApp6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__2; -lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_List_elem___at_Lean_addAliasEntry___spec__16(lean_object*, lean_object*); lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_addCongrTable___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_groundPattern_x3f(lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_isInductivePredicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(lean_object*, lean_object*); +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4; +lean_object* l_Lean_Meta_reduceMatcher_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); lean_object* l_Lean_Expr_toHeadIndex(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2(lean_object*, lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; static size_t l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2; extern lean_object* l_Lean_instInhabitedExpr; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_addCongrTable___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isIte(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__2(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__6; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_mkEMatchEqTheorem(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; -lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__3; lean_object* l_Lean_MessageData_ofExpr(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___boxed(lean_object**); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__4___closed__1; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; -double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_mkGroundPattern(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_constLevels_x21(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +uint8_t l_Lean_Meta_isMatcherAppCore(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3; lean_object* l_Lean_Meta_Grind_registerParent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); +lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); -uint64_t l_Lean_Name_hash___override(lean_object*); -lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +lean_object* lean_get_match_equations_for(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___boxed(lean_object**); -lean_object* l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; lean_object* l_Lean_Meta_Grind_mkENodeCore(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2___boxed(lean_object**); -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_internalize___lambda__2___closed__5; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3; lean_object* l_List_reverse___rarg(lean_object*); -static double l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); -lean_object* lean_array_mk(lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(lean_object*, lean_object*); -lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_internalize___lambda__4___closed__2; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1; size_t lean_usize_add(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___boxed(lean_object**); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__8; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__2; static lean_object* l_Lean_Meta_Grind_internalize___lambda__3___closed__2; -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_Origin_key(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; lean_object* lean_nat_add(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); uint64_t l_Lean_HeadIndex_hash(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2; static lean_object* l_Lean_Meta_Grind_addCongrTable___closed__3; +static lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isLitValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4; lean_object* l_ReaderT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5; static lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at_Lean_Meta_Grind_addCongrTable___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: @@ -527,6 +604,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_G lean_object* x_4; uint64_t x_5; size_t x_6; lean_object* x_7; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); +lean_inc(x_3); x_5 = l_Lean_Meta_Grind_congrHash(x_4, x_3); x_6 = lean_uint64_to_usize(x_5); x_7 = l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2(x_1, x_2, x_6, x_3); @@ -554,6 +632,7 @@ else lean_object* x_11; lean_object* x_12; uint64_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_11 = lean_array_fget(x_3, x_6); x_12 = lean_array_fget(x_4, x_6); +lean_inc(x_11); x_13 = l_Lean_Meta_Grind_congrHash(x_8, x_11); x_14 = lean_uint64_to_usize(x_13); x_15 = 1; @@ -1076,6 +1155,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_ad lean_object* x_5; uint64_t x_6; size_t x_7; size_t x_8; lean_object* x_9; x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); +lean_inc(x_3); x_6 = l_Lean_Meta_Grind_congrHash(x_5, x_3); x_7 = lean_uint64_to_usize(x_6); x_8 = 1; @@ -1083,377 +1163,6 @@ x_9 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___sp return x_9; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 2); -x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -} -static double _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; double x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = 0; -x_3 = l_Float_ofScientific(x_1, x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_12 = lean_ctor_get(x_9, 5); -x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_st_ref_take(x_10, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 3); -lean_inc(x_18); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 0); -lean_dec(x_21); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_17, 3); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) -{ -lean_object* x_25; double x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_25 = lean_ctor_get(x_18, 0); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; -x_27 = 0; -x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; -x_29 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); -lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; -x_31 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_14); -lean_ctor_set(x_31, 2, x_30); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_31); -lean_ctor_set(x_16, 0, x_12); -x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); -lean_ctor_set(x_18, 0, x_32); -x_33 = lean_st_ref_set(x_10, x_17, x_20); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} -} -else -{ -uint64_t x_40; lean_object* x_41; double x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_41 = lean_ctor_get(x_18, 0); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; -x_43 = 0; -x_44 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; -x_45 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); -lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); -x_46 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; -x_47 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_46); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_16, 0, x_12); -x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); -x_49 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); -lean_ctor_set(x_17, 3, x_49); -x_50 = lean_st_ref_set(x_10, x_17, x_20); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_55 = lean_ctor_get(x_17, 0); -x_56 = lean_ctor_get(x_17, 1); -x_57 = lean_ctor_get(x_17, 2); -x_58 = lean_ctor_get(x_17, 4); -x_59 = lean_ctor_get(x_17, 5); -x_60 = lean_ctor_get(x_17, 6); -x_61 = lean_ctor_get(x_17, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_17); -x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_63 = lean_ctor_get(x_18, 0); -lean_inc(x_63); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_64 = x_18; -} else { - lean_dec_ref(x_18); - x_64 = lean_box(0); -} -x_65 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; -x_66 = 0; -x_67 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; -x_68 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); -lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); -x_69 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; -x_70 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_14); -lean_ctor_set(x_70, 2, x_69); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_70); -lean_ctor_set(x_16, 0, x_12); -x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); -if (lean_is_scalar(x_64)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_64; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); -x_73 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_73, 0, x_55); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_57); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_73, 4, x_58); -lean_ctor_set(x_73, 5, x_59); -lean_ctor_set(x_73, 6, x_60); -lean_ctor_set(x_73, 7, x_61); -x_74 = lean_st_ref_set(x_10, x_73, x_20); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); -} -x_77 = lean_box(0); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_75); -return x_78; -} -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_79 = lean_ctor_get(x_16, 1); -lean_inc(x_79); -lean_dec(x_16); -x_80 = lean_ctor_get(x_17, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_17, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_17, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_17, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_17, 5); -lean_inc(x_84); -x_85 = lean_ctor_get(x_17, 6); -lean_inc(x_85); -x_86 = lean_ctor_get(x_17, 7); -lean_inc(x_86); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - lean_ctor_release(x_17, 2); - lean_ctor_release(x_17, 3); - lean_ctor_release(x_17, 4); - lean_ctor_release(x_17, 5); - lean_ctor_release(x_17, 6); - lean_ctor_release(x_17, 7); - x_87 = x_17; -} else { - lean_dec_ref(x_17); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_89 = lean_ctor_get(x_18, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_90 = x_18; -} else { - lean_dec_ref(x_18); - x_90 = lean_box(0); -} -x_91 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1; -x_92 = 0; -x_93 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; -x_94 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); -lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); -x_95 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3; -x_96 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_14); -lean_ctor_set(x_96, 2, x_95); -lean_inc(x_12); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_12); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); -if (lean_is_scalar(x_90)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_90; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); -if (lean_is_scalar(x_87)) { - x_100 = lean_alloc_ctor(0, 8, 0); -} else { - x_100 = x_87; -} -lean_ctor_set(x_100, 0, x_80); -lean_ctor_set(x_100, 1, x_81); -lean_ctor_set(x_100, 2, x_82); -lean_ctor_set(x_100, 3, x_99); -lean_ctor_set(x_100, 4, x_83); -lean_ctor_set(x_100, 5, x_84); -lean_ctor_set(x_100, 6, x_85); -lean_ctor_set(x_100, 7, x_86); -x_101 = lean_st_ref_set(x_10, x_100, x_79); -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_103 = x_101; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -1682,13 +1391,21 @@ return x_4; static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7() { _start: { lean_object* x_1; @@ -1696,11 +1413,11 @@ x_1 = lean_mk_string_unchecked(" = ", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7() { +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -1710,7 +1427,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__2(lean_object { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__4; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_unbox(x_15); @@ -1731,93 +1448,168 @@ uint8_t x_20; x_20 = !lean_is_exclusive(x_14); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_14, 1); x_22 = lean_ctor_get(x_14, 0); lean_dec(x_22); +x_23 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); lean_inc(x_1); -x_23 = l_Lean_MessageData_ofExpr(x_1); -x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; +x_25 = l_Lean_MessageData_ofExpr(x_1); +x_26 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; lean_ctor_set_tag(x_14, 7); -lean_ctor_set(x_14, 1, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_14); -lean_ctor_set(x_26, 1, x_25); -lean_inc(x_2); -x_27 = l_Lean_MessageData_ofExpr(x_2); +lean_ctor_set(x_14, 1, x_25); +lean_ctor_set(x_14, 0, x_26); +x_27 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8; x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 0, x_14); lean_ctor_set(x_28, 1, x_27); -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_24); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Meta_Grind_addCongrTable___lambda__1(x_1, x_2, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); -lean_dec(x_31); -return x_33; +lean_inc(x_2); +x_29 = l_Lean_MessageData_ofExpr(x_2); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_26); +x_32 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Meta_Grind_addCongrTable___lambda__1(x_1, x_2, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +lean_dec(x_33); +return x_35; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_34 = lean_ctor_get(x_14, 1); -lean_inc(x_34); -lean_dec(x_14); -lean_inc(x_1); -x_35 = l_Lean_MessageData_ofExpr(x_1); -x_36 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_35); -x_38 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7; -x_39 = lean_alloc_ctor(7, 2, 0); +uint8_t x_36; +lean_free_object(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_23); +if (x_36 == 0) +{ +return x_23; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_23, 0); +x_38 = lean_ctor_get(x_23, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_23); +x_39 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); -lean_inc(x_2); -x_40 = l_Lean_MessageData_ofExpr(x_2); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_36); -x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = l_Lean_Meta_Grind_addCongrTable___lambda__1(x_1, x_2, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); -lean_dec(x_44); -return x_46; -} +return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -} -static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___closed__1() { -_start: +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_14, 1); +lean_inc(x_40); +lean_dec(x_14); +x_41 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("issues", 6, 6); -return x_1; -} +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +lean_inc(x_1); +x_43 = l_Lean_MessageData_ofExpr(x_1); +x_44 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +x_46 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8; +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_inc(x_2); +x_48 = l_Lean_MessageData_ofExpr(x_2); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_44); +x_51 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_50, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_42); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_Meta_Grind_addCongrTable___lambda__1(x_1, x_2, x_52, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_53); +lean_dec(x_52); +return x_54; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_1); +x_55 = lean_ctor_get(x_41, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_41, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_57 = x_41; +} else { + lean_dec_ref(x_41); + x_57 = lean_box(0); +} +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(1, 2, 0); +} else { + x_58 = x_57; +} +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_56); +return x_58; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +} +static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("issues", 6, 6); +return x_1; +} } static lean_object* _init_l_Lean_Meta_Grind_addCongrTable___closed__2() { _start: @@ -1905,7 +1697,7 @@ lean_inc(x_1); x_16 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_13, x_15, x_1); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -1932,7 +1724,7 @@ x_24 = lean_ctor_get(x_18, 4); lean_inc(x_24); x_25 = lean_ctor_get(x_18, 5); lean_inc(x_25); -x_26 = lean_ctor_get_uint8(x_18, sizeof(void*)*14); +x_26 = lean_ctor_get_uint8(x_18, sizeof(void*)*19); x_27 = lean_ctor_get(x_18, 6); lean_inc(x_27); x_28 = lean_ctor_get(x_18, 7); @@ -1949,272 +1741,384 @@ x_33 = lean_ctor_get(x_18, 12); lean_inc(x_33); x_34 = lean_ctor_get(x_18, 13); lean_inc(x_34); -x_35 = lean_box(0); +x_35 = lean_ctor_get(x_18, 14); +lean_inc(x_35); +x_36 = lean_ctor_get(x_18, 15); +lean_inc(x_36); +x_37 = lean_ctor_get(x_18, 16); +lean_inc(x_37); +x_38 = lean_ctor_get(x_18, 17); +lean_inc(x_38); +x_39 = lean_ctor_get(x_18, 18); +lean_inc(x_39); +x_40 = lean_box(0); lean_inc(x_18); -x_36 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_18, x_23, x_1, x_35); -x_37 = !lean_is_exclusive(x_18); -if (x_37 == 0) +x_41 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_18, x_23, x_1, x_40); +x_42 = !lean_is_exclusive(x_18); +if (x_42 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_38 = lean_ctor_get(x_18, 13); -lean_dec(x_38); -x_39 = lean_ctor_get(x_18, 12); -lean_dec(x_39); -x_40 = lean_ctor_get(x_18, 11); -lean_dec(x_40); -x_41 = lean_ctor_get(x_18, 10); -lean_dec(x_41); -x_42 = lean_ctor_get(x_18, 9); -lean_dec(x_42); -x_43 = lean_ctor_get(x_18, 8); +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_43 = lean_ctor_get(x_18, 18); lean_dec(x_43); -x_44 = lean_ctor_get(x_18, 7); +x_44 = lean_ctor_get(x_18, 17); lean_dec(x_44); -x_45 = lean_ctor_get(x_18, 6); +x_45 = lean_ctor_get(x_18, 16); lean_dec(x_45); -x_46 = lean_ctor_get(x_18, 5); +x_46 = lean_ctor_get(x_18, 15); lean_dec(x_46); -x_47 = lean_ctor_get(x_18, 4); +x_47 = lean_ctor_get(x_18, 14); lean_dec(x_47); -x_48 = lean_ctor_get(x_18, 3); +x_48 = lean_ctor_get(x_18, 13); lean_dec(x_48); -x_49 = lean_ctor_get(x_18, 2); +x_49 = lean_ctor_get(x_18, 12); lean_dec(x_49); -x_50 = lean_ctor_get(x_18, 1); +x_50 = lean_ctor_get(x_18, 11); lean_dec(x_50); -x_51 = lean_ctor_get(x_18, 0); +x_51 = lean_ctor_get(x_18, 10); lean_dec(x_51); -lean_ctor_set(x_18, 3, x_36); -x_52 = lean_st_ref_set(x_2, x_18, x_19); +x_52 = lean_ctor_get(x_18, 9); +lean_dec(x_52); +x_53 = lean_ctor_get(x_18, 8); +lean_dec(x_53); +x_54 = lean_ctor_get(x_18, 7); +lean_dec(x_54); +x_55 = lean_ctor_get(x_18, 6); +lean_dec(x_55); +x_56 = lean_ctor_get(x_18, 5); +lean_dec(x_56); +x_57 = lean_ctor_get(x_18, 4); +lean_dec(x_57); +x_58 = lean_ctor_get(x_18, 3); +lean_dec(x_58); +x_59 = lean_ctor_get(x_18, 2); +lean_dec(x_59); +x_60 = lean_ctor_get(x_18, 1); +lean_dec(x_60); +x_61 = lean_ctor_get(x_18, 0); +lean_dec(x_61); +lean_ctor_set(x_18, 3, x_41); +x_62 = lean_st_ref_set(x_2, x_18, x_19); lean_dec(x_2); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +x_63 = !lean_is_exclusive(x_62); +if (x_63 == 0) { -lean_object* x_54; -x_54 = lean_ctor_get(x_52, 0); -lean_dec(x_54); -lean_ctor_set(x_52, 0, x_35); -return x_52; +lean_object* x_64; +x_64 = lean_ctor_get(x_62, 0); +lean_dec(x_64); +lean_ctor_set(x_62, 0, x_40); +return x_62; } else { -lean_object* x_55; lean_object* x_56; -x_55 = lean_ctor_get(x_52, 1); -lean_inc(x_55); -lean_dec(x_52); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_35); -lean_ctor_set(x_56, 1, x_55); -return x_56; +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_62, 1); +lean_inc(x_65); +lean_dec(x_62); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_40); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_dec(x_18); -x_57 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_57, 0, x_20); -lean_ctor_set(x_57, 1, x_21); -lean_ctor_set(x_57, 2, x_22); -lean_ctor_set(x_57, 3, x_36); -lean_ctor_set(x_57, 4, x_24); -lean_ctor_set(x_57, 5, x_25); -lean_ctor_set(x_57, 6, x_27); -lean_ctor_set(x_57, 7, x_28); -lean_ctor_set(x_57, 8, x_29); -lean_ctor_set(x_57, 9, x_30); -lean_ctor_set(x_57, 10, x_31); -lean_ctor_set(x_57, 11, x_32); -lean_ctor_set(x_57, 12, x_33); -lean_ctor_set(x_57, 13, x_34); -lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_26); -x_58 = lean_st_ref_set(x_2, x_57, x_19); +x_67 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_67, 0, x_20); +lean_ctor_set(x_67, 1, x_21); +lean_ctor_set(x_67, 2, x_22); +lean_ctor_set(x_67, 3, x_41); +lean_ctor_set(x_67, 4, x_24); +lean_ctor_set(x_67, 5, x_25); +lean_ctor_set(x_67, 6, x_27); +lean_ctor_set(x_67, 7, x_28); +lean_ctor_set(x_67, 8, x_29); +lean_ctor_set(x_67, 9, x_30); +lean_ctor_set(x_67, 10, x_31); +lean_ctor_set(x_67, 11, x_32); +lean_ctor_set(x_67, 12, x_33); +lean_ctor_set(x_67, 13, x_34); +lean_ctor_set(x_67, 14, x_35); +lean_ctor_set(x_67, 15, x_36); +lean_ctor_set(x_67, 16, x_37); +lean_ctor_set(x_67, 17, x_38); +lean_ctor_set(x_67, 18, x_39); +lean_ctor_set_uint8(x_67, sizeof(void*)*19, x_26); +x_68 = lean_st_ref_set(x_2, x_67, x_19); lean_dec(x_2); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - lean_ctor_release(x_58, 1); - x_60 = x_58; +x_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_58); - x_60 = lean_box(0); + lean_dec_ref(x_68); + x_70 = lean_box(0); } -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_70)) { + x_71 = lean_alloc_ctor(0, 2, 0); } else { - x_61 = x_60; + x_71 = x_70; } -lean_ctor_set(x_61, 0, x_35); -lean_ctor_set(x_61, 1, x_59); -return x_61; +lean_ctor_set(x_71, 0, x_40); +lean_ctor_set(x_71, 1, x_69); +return x_71; } } else { -lean_object* x_62; uint8_t x_63; -x_62 = lean_ctor_get(x_16, 0); -lean_inc(x_62); +lean_object* x_72; uint8_t x_73; +x_72 = lean_ctor_get(x_16, 0); +lean_inc(x_72); lean_dec(x_16); -x_63 = !lean_is_exclusive(x_62); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_64 = lean_ctor_get(x_62, 0); -x_65 = lean_ctor_get(x_62, 1); -lean_dec(x_65); -x_66 = l_Lean_Expr_getAppFn(x_1); -x_67 = l_Lean_Expr_getAppFn(x_64); -x_68 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_66, x_67); -if (x_68 == 0) +x_73 = !lean_is_exclusive(x_72); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_74 = lean_ctor_get(x_72, 0); +x_75 = lean_ctor_get(x_72, 1); +lean_dec(x_75); +x_76 = l_Lean_Expr_getAppFn(x_1); +x_77 = l_Lean_Expr_getAppFn(x_74); +x_78 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_76, x_77); +if (x_78 == 0) { -lean_object* x_69; +lean_object* x_79; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_69 = l_Lean_Meta_Grind_hasSameType(x_66, x_67, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_69) == 0) +x_79 = l_Lean_Meta_Grind_hasSameType(x_76, x_77, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_70; uint8_t x_71; -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_unbox(x_70); -lean_dec(x_70); -if (x_71 == 0) +lean_object* x_80; uint8_t x_81; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_unbox(x_80); +lean_dec(x_80); +if (x_81 == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; -x_72 = lean_ctor_get(x_69, 1); -lean_inc(x_72); -lean_dec(x_69); -x_73 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_74 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_73, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_72); -x_75 = !lean_is_exclusive(x_74); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; -x_76 = lean_ctor_get(x_74, 0); -x_77 = lean_ctor_get(x_74, 1); -x_78 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_79 = lean_unbox(x_76); -lean_dec(x_76); -if (x_79 == 0) +lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_82 = lean_ctor_get(x_79, 1); +lean_inc(x_82); +lean_dec(x_79); +x_83 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_84 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_83, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_82); +x_85 = !lean_is_exclusive(x_84); +if (x_85 == 0) { -lean_object* x_80; lean_object* x_81; -lean_free_object(x_74); -lean_free_object(x_62); -lean_dec(x_64); +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_84, 0); +x_87 = lean_ctor_get(x_84, 1); +x_88 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_89 = lean_unbox(x_86); +lean_dec(x_86); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; +lean_free_object(x_84); +lean_free_object(x_72); +lean_dec(x_74); lean_free_object(x_11); lean_dec(x_1); -x_80 = lean_box(0); -x_81 = lean_apply_10(x_78, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_77); -return x_81; +x_90 = lean_box(0); +x_91 = lean_apply_10(x_88, x_90, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_87); +return x_91; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_82 = l_Lean_indentExpr(x_1); -x_83 = l_Lean_Meta_Grind_addCongrTable___closed__5; -lean_ctor_set_tag(x_74, 7); -lean_ctor_set(x_74, 1, x_82); -lean_ctor_set(x_74, 0, x_83); -x_84 = l_Lean_Meta_Grind_addCongrTable___closed__7; -lean_ctor_set_tag(x_62, 7); -lean_ctor_set(x_62, 1, x_84); -lean_ctor_set(x_62, 0, x_74); -x_85 = l_Lean_indentExpr(x_64); +lean_object* x_92; +x_92 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_87); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +lean_dec(x_92); +x_94 = l_Lean_indentExpr(x_1); +x_95 = l_Lean_Meta_Grind_addCongrTable___closed__5; +lean_ctor_set_tag(x_84, 7); +lean_ctor_set(x_84, 1, x_94); +lean_ctor_set(x_84, 0, x_95); +x_96 = l_Lean_Meta_Grind_addCongrTable___closed__7; +lean_ctor_set_tag(x_72, 7); +lean_ctor_set(x_72, 1, x_96); +lean_ctor_set(x_72, 0, x_84); +x_97 = l_Lean_indentExpr(x_74); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_85); -lean_ctor_set(x_11, 0, x_62); -x_86 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_11); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_73, x_87, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_77); -x_89 = lean_ctor_get(x_88, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_88, 1); -lean_inc(x_90); -lean_dec(x_88); -x_91 = lean_apply_10(x_78, x_89, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_90); -return x_91; -} +lean_ctor_set(x_11, 1, x_97); +lean_ctor_set(x_11, 0, x_72); +x_98 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_99 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_99, 0, x_11); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_83, x_99, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_93); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_100, 1); +lean_inc(x_102); +lean_dec(x_100); +x_103 = lean_apply_10(x_88, x_101, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_102); +return x_103; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_92 = lean_ctor_get(x_74, 0); -x_93 = lean_ctor_get(x_74, 1); -lean_inc(x_93); -lean_inc(x_92); +uint8_t x_104; +lean_free_object(x_84); +lean_free_object(x_72); lean_dec(x_74); -x_94 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_95 = lean_unbox(x_92); +lean_free_object(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_104 = !lean_is_exclusive(x_92); +if (x_104 == 0) +{ +return x_92; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_92, 0); +x_106 = lean_ctor_get(x_92, 1); +lean_inc(x_106); +lean_inc(x_105); lean_dec(x_92); -if (x_95 == 0) +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +return x_107; +} +} +} +} +else { -lean_object* x_96; lean_object* x_97; -lean_free_object(x_62); -lean_dec(x_64); +lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; +x_108 = lean_ctor_get(x_84, 0); +x_109 = lean_ctor_get(x_84, 1); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_84); +x_110 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_111 = lean_unbox(x_108); +lean_dec(x_108); +if (x_111 == 0) +{ +lean_object* x_112; lean_object* x_113; +lean_free_object(x_72); +lean_dec(x_74); lean_free_object(x_11); lean_dec(x_1); -x_96 = lean_box(0); -x_97 = lean_apply_10(x_94, x_96, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_93); -return x_97; +x_112 = lean_box(0); +x_113 = lean_apply_10(x_110, x_112, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_109); +return x_113; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_98 = l_Lean_indentExpr(x_1); -x_99 = l_Lean_Meta_Grind_addCongrTable___closed__5; -x_100 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_98); -x_101 = l_Lean_Meta_Grind_addCongrTable___closed__7; -lean_ctor_set_tag(x_62, 7); -lean_ctor_set(x_62, 1, x_101); -lean_ctor_set(x_62, 0, x_100); -x_102 = l_Lean_indentExpr(x_64); +lean_object* x_114; +x_114 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_109); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = l_Lean_indentExpr(x_1); +x_117 = l_Lean_Meta_Grind_addCongrTable___closed__5; +x_118 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_116); +x_119 = l_Lean_Meta_Grind_addCongrTable___closed__7; +lean_ctor_set_tag(x_72, 7); +lean_ctor_set(x_72, 1, x_119); +lean_ctor_set(x_72, 0, x_118); +x_120 = l_Lean_indentExpr(x_74); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_102); -lean_ctor_set(x_11, 0, x_62); -x_103 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_104 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_104, 0, x_11); -lean_ctor_set(x_104, 1, x_103); -x_105 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_73, x_104, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_93); -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -lean_dec(x_105); -x_108 = lean_apply_10(x_94, x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_107); -return x_108; +lean_ctor_set(x_11, 1, x_120); +lean_ctor_set(x_11, 0, x_72); +x_121 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_122 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_122, 0, x_11); +lean_ctor_set(x_122, 1, x_121); +x_123 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_83, x_122, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_115); +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +lean_dec(x_123); +x_126 = lean_apply_10(x_110, x_124, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_125); +return x_126; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +lean_free_object(x_72); +lean_dec(x_74); +lean_free_object(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_127 = lean_ctor_get(x_114, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_114, 1); +lean_inc(x_128); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_129 = x_114; +} else { + lean_dec_ref(x_114); + x_129 = lean_box(0); +} +if (lean_is_scalar(x_129)) { + x_130 = lean_alloc_ctor(1, 2, 0); +} else { + x_130 = x_129; +} +lean_ctor_set(x_130, 0, x_127); +lean_ctor_set(x_130, 1, x_128); +return x_130; +} } } } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_free_object(x_62); +lean_object* x_131; lean_object* x_132; lean_object* x_133; +lean_free_object(x_72); lean_free_object(x_11); -x_109 = lean_ctor_get(x_69, 1); -lean_inc(x_109); -lean_dec(x_69); -x_110 = lean_box(0); -x_111 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_64, x_110, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_109); +x_131 = lean_ctor_get(x_79, 1); +lean_inc(x_131); +lean_dec(x_79); +x_132 = lean_box(0); +x_133 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_74, x_132, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_131); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_111; +return x_133; } } else { -uint8_t x_112; -lean_free_object(x_62); -lean_dec(x_64); +uint8_t x_134; +lean_free_object(x_72); +lean_dec(x_74); lean_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -2225,155 +2129,199 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_112 = !lean_is_exclusive(x_69); -if (x_112 == 0) +x_134 = !lean_is_exclusive(x_79); +if (x_134 == 0) { -return x_69; +return x_79; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_69, 0); -x_114 = lean_ctor_get(x_69, 1); -lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_69); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -return x_115; +lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_135 = lean_ctor_get(x_79, 0); +x_136 = lean_ctor_get(x_79, 1); +lean_inc(x_136); +lean_inc(x_135); +lean_dec(x_79); +x_137 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +return x_137; } } } else { -lean_object* x_116; lean_object* x_117; -lean_dec(x_67); -lean_dec(x_66); -lean_free_object(x_62); +lean_object* x_138; lean_object* x_139; +lean_dec(x_77); +lean_dec(x_76); +lean_free_object(x_72); lean_free_object(x_11); -x_116 = lean_box(0); -x_117 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_64, x_116, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_138 = lean_box(0); +x_139 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_74, x_138, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_117; +return x_139; } } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; -x_118 = lean_ctor_get(x_62, 0); -lean_inc(x_118); -lean_dec(x_62); -x_119 = l_Lean_Expr_getAppFn(x_1); -x_120 = l_Lean_Expr_getAppFn(x_118); -x_121 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_119, x_120); -if (x_121 == 0) +lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; +x_140 = lean_ctor_get(x_72, 0); +lean_inc(x_140); +lean_dec(x_72); +x_141 = l_Lean_Expr_getAppFn(x_1); +x_142 = l_Lean_Expr_getAppFn(x_140); +x_143 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_141, x_142); +if (x_143 == 0) { -lean_object* x_122; +lean_object* x_144; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_122 = l_Lean_Meta_Grind_hasSameType(x_119, x_120, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_122) == 0) +x_144 = l_Lean_Meta_Grind_hasSameType(x_141, x_142, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_144) == 0) { -lean_object* x_123; uint8_t x_124; -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_unbox(x_123); -lean_dec(x_123); -if (x_124 == 0) +lean_object* x_145; uint8_t x_146; +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_unbox(x_145); +lean_dec(x_145); +if (x_146 == 0) { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_125 = lean_ctor_get(x_122, 1); -lean_inc(x_125); -lean_dec(x_122); -x_126 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_127 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_126, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_125); -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_ctor_get(x_127, 1); -lean_inc(x_129); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - lean_ctor_release(x_127, 1); - x_130 = x_127; +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; +x_147 = lean_ctor_get(x_144, 1); +lean_inc(x_147); +lean_dec(x_144); +x_148 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_149 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_148, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_147); +x_150 = lean_ctor_get(x_149, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_149, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_149)) { + lean_ctor_release(x_149, 0); + lean_ctor_release(x_149, 1); + x_152 = x_149; } else { - lean_dec_ref(x_127); - x_130 = lean_box(0); + lean_dec_ref(x_149); + x_152 = lean_box(0); } -x_131 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_132 = lean_unbox(x_128); -lean_dec(x_128); -if (x_132 == 0) +x_153 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_154 = lean_unbox(x_150); +lean_dec(x_150); +if (x_154 == 0) { -lean_object* x_133; lean_object* x_134; -lean_dec(x_130); -lean_dec(x_118); +lean_object* x_155; lean_object* x_156; +lean_dec(x_152); +lean_dec(x_140); lean_free_object(x_11); lean_dec(x_1); -x_133 = lean_box(0); -x_134 = lean_apply_10(x_131, x_133, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_129); -return x_134; +x_155 = lean_box(0); +x_156 = lean_apply_10(x_153, x_155, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_151); +return x_156; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_135 = l_Lean_indentExpr(x_1); -x_136 = l_Lean_Meta_Grind_addCongrTable___closed__5; -if (lean_is_scalar(x_130)) { - x_137 = lean_alloc_ctor(7, 2, 0); +lean_object* x_157; +x_157 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_151); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_158 = lean_ctor_get(x_157, 1); +lean_inc(x_158); +lean_dec(x_157); +x_159 = l_Lean_indentExpr(x_1); +x_160 = l_Lean_Meta_Grind_addCongrTable___closed__5; +if (lean_is_scalar(x_152)) { + x_161 = lean_alloc_ctor(7, 2, 0); } else { - x_137 = x_130; - lean_ctor_set_tag(x_137, 7); -} -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_135); -x_138 = l_Lean_Meta_Grind_addCongrTable___closed__7; -x_139 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_139, 0, x_137); -lean_ctor_set(x_139, 1, x_138); -x_140 = l_Lean_indentExpr(x_118); + x_161 = x_152; + lean_ctor_set_tag(x_161, 7); +} +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_159); +x_162 = l_Lean_Meta_Grind_addCongrTable___closed__7; +x_163 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_163, 0, x_161); +lean_ctor_set(x_163, 1, x_162); +x_164 = l_Lean_indentExpr(x_140); lean_ctor_set_tag(x_11, 7); -lean_ctor_set(x_11, 1, x_140); -lean_ctor_set(x_11, 0, x_139); -x_141 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_142 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_142, 0, x_11); -lean_ctor_set(x_142, 1, x_141); -x_143 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_126, x_142, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_129); -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_143, 1); -lean_inc(x_145); -lean_dec(x_143); -x_146 = lean_apply_10(x_131, x_144, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_145); -return x_146; +lean_ctor_set(x_11, 1, x_164); +lean_ctor_set(x_11, 0, x_163); +x_165 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_166 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_166, 0, x_11); +lean_ctor_set(x_166, 1, x_165); +x_167 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_148, x_166, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_158); +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_167, 1); +lean_inc(x_169); +lean_dec(x_167); +x_170 = lean_apply_10(x_153, x_168, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_169); +return x_170; +} +else +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +lean_dec(x_152); +lean_dec(x_140); +lean_free_object(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_171 = lean_ctor_get(x_157, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_157, 1); +lean_inc(x_172); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_173 = x_157; +} else { + lean_dec_ref(x_157); + x_173 = lean_box(0); +} +if (lean_is_scalar(x_173)) { + x_174 = lean_alloc_ctor(1, 2, 0); +} else { + x_174 = x_173; +} +lean_ctor_set(x_174, 0, x_171); +lean_ctor_set(x_174, 1, x_172); +return x_174; +} } } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_free_object(x_11); -x_147 = lean_ctor_get(x_122, 1); -lean_inc(x_147); -lean_dec(x_122); -x_148 = lean_box(0); -x_149 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_118, x_148, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_147); +x_175 = lean_ctor_get(x_144, 1); +lean_inc(x_175); +lean_dec(x_144); +x_176 = lean_box(0); +x_177 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_140, x_176, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_175); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_149; +return x_177; } } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -lean_dec(x_118); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +lean_dec(x_140); lean_free_object(x_11); lean_dec(x_9); lean_dec(x_8); @@ -2384,60 +2332,60 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_150 = lean_ctor_get(x_122, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_122, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_152 = x_122; +x_178 = lean_ctor_get(x_144, 0); +lean_inc(x_178); +x_179 = lean_ctor_get(x_144, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_180 = x_144; } else { - lean_dec_ref(x_122); - x_152 = lean_box(0); + lean_dec_ref(x_144); + x_180 = lean_box(0); } -if (lean_is_scalar(x_152)) { - x_153 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_180)) { + x_181 = lean_alloc_ctor(1, 2, 0); } else { - x_153 = x_152; + x_181 = x_180; } -lean_ctor_set(x_153, 0, x_150); -lean_ctor_set(x_153, 1, x_151); -return x_153; +lean_ctor_set(x_181, 0, x_178); +lean_ctor_set(x_181, 1, x_179); +return x_181; } } else { -lean_object* x_154; lean_object* x_155; -lean_dec(x_120); -lean_dec(x_119); +lean_object* x_182; lean_object* x_183; +lean_dec(x_142); +lean_dec(x_141); lean_free_object(x_11); -x_154 = lean_box(0); -x_155 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_118, x_154, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_182 = lean_box(0); +x_183 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_140, x_182, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_155; +return x_183; } } } } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_156 = lean_ctor_get(x_11, 0); -x_157 = lean_ctor_get(x_11, 1); -lean_inc(x_157); -lean_inc(x_156); +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_184 = lean_ctor_get(x_11, 0); +x_185 = lean_ctor_get(x_11, 1); +lean_inc(x_185); +lean_inc(x_184); lean_dec(x_11); -x_158 = lean_ctor_get(x_156, 3); -lean_inc(x_158); +x_186 = lean_ctor_get(x_184, 3); +lean_inc(x_186); lean_inc(x_1); -x_159 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_156, x_158, x_1); -if (lean_obj_tag(x_159) == 0) +x_187 = l_Lean_PersistentHashMap_findEntry_x3f___at_Lean_Meta_Grind_addCongrTable___spec__1(x_184, x_186, x_1); +if (lean_obj_tag(x_187) == 0) { -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; uint8_t x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2445,234 +2393,298 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_160 = lean_st_ref_take(x_2, x_157); -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_160, 1); -lean_inc(x_162); -lean_dec(x_160); -x_163 = lean_ctor_get(x_161, 0); -lean_inc(x_163); -x_164 = lean_ctor_get(x_161, 1); -lean_inc(x_164); -x_165 = lean_ctor_get(x_161, 2); -lean_inc(x_165); -x_166 = lean_ctor_get(x_161, 3); -lean_inc(x_166); -x_167 = lean_ctor_get(x_161, 4); -lean_inc(x_167); -x_168 = lean_ctor_get(x_161, 5); -lean_inc(x_168); -x_169 = lean_ctor_get_uint8(x_161, sizeof(void*)*14); -x_170 = lean_ctor_get(x_161, 6); -lean_inc(x_170); -x_171 = lean_ctor_get(x_161, 7); -lean_inc(x_171); -x_172 = lean_ctor_get(x_161, 8); -lean_inc(x_172); -x_173 = lean_ctor_get(x_161, 9); -lean_inc(x_173); -x_174 = lean_ctor_get(x_161, 10); -lean_inc(x_174); -x_175 = lean_ctor_get(x_161, 11); -lean_inc(x_175); -x_176 = lean_ctor_get(x_161, 12); -lean_inc(x_176); -x_177 = lean_ctor_get(x_161, 13); -lean_inc(x_177); -x_178 = lean_box(0); -lean_inc(x_161); -x_179 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_161, x_166, x_1, x_178); -if (lean_is_exclusive(x_161)) { - lean_ctor_release(x_161, 0); - lean_ctor_release(x_161, 1); - lean_ctor_release(x_161, 2); - lean_ctor_release(x_161, 3); - lean_ctor_release(x_161, 4); - lean_ctor_release(x_161, 5); - lean_ctor_release(x_161, 6); - lean_ctor_release(x_161, 7); - lean_ctor_release(x_161, 8); - lean_ctor_release(x_161, 9); - lean_ctor_release(x_161, 10); - lean_ctor_release(x_161, 11); - lean_ctor_release(x_161, 12); - lean_ctor_release(x_161, 13); - x_180 = x_161; +x_188 = lean_st_ref_take(x_2, x_185); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +lean_dec(x_188); +x_191 = lean_ctor_get(x_189, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_189, 1); +lean_inc(x_192); +x_193 = lean_ctor_get(x_189, 2); +lean_inc(x_193); +x_194 = lean_ctor_get(x_189, 3); +lean_inc(x_194); +x_195 = lean_ctor_get(x_189, 4); +lean_inc(x_195); +x_196 = lean_ctor_get(x_189, 5); +lean_inc(x_196); +x_197 = lean_ctor_get_uint8(x_189, sizeof(void*)*19); +x_198 = lean_ctor_get(x_189, 6); +lean_inc(x_198); +x_199 = lean_ctor_get(x_189, 7); +lean_inc(x_199); +x_200 = lean_ctor_get(x_189, 8); +lean_inc(x_200); +x_201 = lean_ctor_get(x_189, 9); +lean_inc(x_201); +x_202 = lean_ctor_get(x_189, 10); +lean_inc(x_202); +x_203 = lean_ctor_get(x_189, 11); +lean_inc(x_203); +x_204 = lean_ctor_get(x_189, 12); +lean_inc(x_204); +x_205 = lean_ctor_get(x_189, 13); +lean_inc(x_205); +x_206 = lean_ctor_get(x_189, 14); +lean_inc(x_206); +x_207 = lean_ctor_get(x_189, 15); +lean_inc(x_207); +x_208 = lean_ctor_get(x_189, 16); +lean_inc(x_208); +x_209 = lean_ctor_get(x_189, 17); +lean_inc(x_209); +x_210 = lean_ctor_get(x_189, 18); +lean_inc(x_210); +x_211 = lean_box(0); +lean_inc(x_189); +x_212 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_addCongrTable___spec__4(x_189, x_194, x_1, x_211); +if (lean_is_exclusive(x_189)) { + lean_ctor_release(x_189, 0); + lean_ctor_release(x_189, 1); + lean_ctor_release(x_189, 2); + lean_ctor_release(x_189, 3); + lean_ctor_release(x_189, 4); + lean_ctor_release(x_189, 5); + lean_ctor_release(x_189, 6); + lean_ctor_release(x_189, 7); + lean_ctor_release(x_189, 8); + lean_ctor_release(x_189, 9); + lean_ctor_release(x_189, 10); + lean_ctor_release(x_189, 11); + lean_ctor_release(x_189, 12); + lean_ctor_release(x_189, 13); + lean_ctor_release(x_189, 14); + lean_ctor_release(x_189, 15); + lean_ctor_release(x_189, 16); + lean_ctor_release(x_189, 17); + lean_ctor_release(x_189, 18); + x_213 = x_189; } else { - lean_dec_ref(x_161); - x_180 = lean_box(0); + lean_dec_ref(x_189); + x_213 = lean_box(0); } -if (lean_is_scalar(x_180)) { - x_181 = lean_alloc_ctor(0, 14, 1); +if (lean_is_scalar(x_213)) { + x_214 = lean_alloc_ctor(0, 19, 1); } else { - x_181 = x_180; -} -lean_ctor_set(x_181, 0, x_163); -lean_ctor_set(x_181, 1, x_164); -lean_ctor_set(x_181, 2, x_165); -lean_ctor_set(x_181, 3, x_179); -lean_ctor_set(x_181, 4, x_167); -lean_ctor_set(x_181, 5, x_168); -lean_ctor_set(x_181, 6, x_170); -lean_ctor_set(x_181, 7, x_171); -lean_ctor_set(x_181, 8, x_172); -lean_ctor_set(x_181, 9, x_173); -lean_ctor_set(x_181, 10, x_174); -lean_ctor_set(x_181, 11, x_175); -lean_ctor_set(x_181, 12, x_176); -lean_ctor_set(x_181, 13, x_177); -lean_ctor_set_uint8(x_181, sizeof(void*)*14, x_169); -x_182 = lean_st_ref_set(x_2, x_181, x_162); -lean_dec(x_2); -x_183 = lean_ctor_get(x_182, 1); -lean_inc(x_183); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_184 = x_182; + x_214 = x_213; +} +lean_ctor_set(x_214, 0, x_191); +lean_ctor_set(x_214, 1, x_192); +lean_ctor_set(x_214, 2, x_193); +lean_ctor_set(x_214, 3, x_212); +lean_ctor_set(x_214, 4, x_195); +lean_ctor_set(x_214, 5, x_196); +lean_ctor_set(x_214, 6, x_198); +lean_ctor_set(x_214, 7, x_199); +lean_ctor_set(x_214, 8, x_200); +lean_ctor_set(x_214, 9, x_201); +lean_ctor_set(x_214, 10, x_202); +lean_ctor_set(x_214, 11, x_203); +lean_ctor_set(x_214, 12, x_204); +lean_ctor_set(x_214, 13, x_205); +lean_ctor_set(x_214, 14, x_206); +lean_ctor_set(x_214, 15, x_207); +lean_ctor_set(x_214, 16, x_208); +lean_ctor_set(x_214, 17, x_209); +lean_ctor_set(x_214, 18, x_210); +lean_ctor_set_uint8(x_214, sizeof(void*)*19, x_197); +x_215 = lean_st_ref_set(x_2, x_214, x_190); +lean_dec(x_2); +x_216 = lean_ctor_get(x_215, 1); +lean_inc(x_216); +if (lean_is_exclusive(x_215)) { + lean_ctor_release(x_215, 0); + lean_ctor_release(x_215, 1); + x_217 = x_215; } else { - lean_dec_ref(x_182); - x_184 = lean_box(0); + lean_dec_ref(x_215); + x_217 = lean_box(0); } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_217)) { + x_218 = lean_alloc_ctor(0, 2, 0); } else { - x_185 = x_184; + x_218 = x_217; } -lean_ctor_set(x_185, 0, x_178); -lean_ctor_set(x_185, 1, x_183); -return x_185; +lean_ctor_set(x_218, 0, x_211); +lean_ctor_set(x_218, 1, x_216); +return x_218; } else { -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; -x_186 = lean_ctor_get(x_159, 0); -lean_inc(x_186); -lean_dec(x_159); -x_187 = lean_ctor_get(x_186, 0); -lean_inc(x_187); -if (lean_is_exclusive(x_186)) { - lean_ctor_release(x_186, 0); - lean_ctor_release(x_186, 1); - x_188 = x_186; +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; uint8_t x_224; +x_219 = lean_ctor_get(x_187, 0); +lean_inc(x_219); +lean_dec(x_187); +x_220 = lean_ctor_get(x_219, 0); +lean_inc(x_220); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_221 = x_219; } else { - lean_dec_ref(x_186); - x_188 = lean_box(0); + lean_dec_ref(x_219); + x_221 = lean_box(0); } -x_189 = l_Lean_Expr_getAppFn(x_1); -x_190 = l_Lean_Expr_getAppFn(x_187); -x_191 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_189, x_190); -if (x_191 == 0) +x_222 = l_Lean_Expr_getAppFn(x_1); +x_223 = l_Lean_Expr_getAppFn(x_220); +x_224 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_222, x_223); +if (x_224 == 0) { -lean_object* x_192; +lean_object* x_225; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_192 = l_Lean_Meta_Grind_hasSameType(x_189, x_190, x_6, x_7, x_8, x_9, x_157); -if (lean_obj_tag(x_192) == 0) +x_225 = l_Lean_Meta_Grind_hasSameType(x_222, x_223, x_6, x_7, x_8, x_9, x_185); +if (lean_obj_tag(x_225) == 0) { -lean_object* x_193; uint8_t x_194; -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_unbox(x_193); -lean_dec(x_193); -if (x_194 == 0) +lean_object* x_226; uint8_t x_227; +x_226 = lean_ctor_get(x_225, 0); +lean_inc(x_226); +x_227 = lean_unbox(x_226); +lean_dec(x_226); +if (x_227 == 0) { -lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; uint8_t x_202; -x_195 = lean_ctor_get(x_192, 1); -lean_inc(x_195); -lean_dec(x_192); -x_196 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_197 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_196, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_195); -x_198 = lean_ctor_get(x_197, 0); -lean_inc(x_198); -x_199 = lean_ctor_get(x_197, 1); -lean_inc(x_199); -if (lean_is_exclusive(x_197)) { - lean_ctor_release(x_197, 0); - lean_ctor_release(x_197, 1); - x_200 = x_197; +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; uint8_t x_235; +x_228 = lean_ctor_get(x_225, 1); +lean_inc(x_228); +lean_dec(x_225); +x_229 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_230 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_229, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_228); +x_231 = lean_ctor_get(x_230, 0); +lean_inc(x_231); +x_232 = lean_ctor_get(x_230, 1); +lean_inc(x_232); +if (lean_is_exclusive(x_230)) { + lean_ctor_release(x_230, 0); + lean_ctor_release(x_230, 1); + x_233 = x_230; } else { - lean_dec_ref(x_197); - x_200 = lean_box(0); + lean_dec_ref(x_230); + x_233 = lean_box(0); } -x_201 = l_Lean_Meta_Grind_addCongrTable___closed__3; -x_202 = lean_unbox(x_198); -lean_dec(x_198); -if (x_202 == 0) +x_234 = l_Lean_Meta_Grind_addCongrTable___closed__3; +x_235 = lean_unbox(x_231); +lean_dec(x_231); +if (x_235 == 0) { -lean_object* x_203; lean_object* x_204; -lean_dec(x_200); -lean_dec(x_188); -lean_dec(x_187); +lean_object* x_236; lean_object* x_237; +lean_dec(x_233); +lean_dec(x_221); +lean_dec(x_220); lean_dec(x_1); -x_203 = lean_box(0); -x_204 = lean_apply_10(x_201, x_203, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_199); -return x_204; +x_236 = lean_box(0); +x_237 = lean_apply_10(x_234, x_236, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_232); +return x_237; } else { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; -x_205 = l_Lean_indentExpr(x_1); -x_206 = l_Lean_Meta_Grind_addCongrTable___closed__5; -if (lean_is_scalar(x_200)) { - x_207 = lean_alloc_ctor(7, 2, 0); +lean_object* x_238; +x_238 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_232); +if (lean_obj_tag(x_238) == 0) +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_239 = lean_ctor_get(x_238, 1); +lean_inc(x_239); +lean_dec(x_238); +x_240 = l_Lean_indentExpr(x_1); +x_241 = l_Lean_Meta_Grind_addCongrTable___closed__5; +if (lean_is_scalar(x_233)) { + x_242 = lean_alloc_ctor(7, 2, 0); } else { - x_207 = x_200; - lean_ctor_set_tag(x_207, 7); -} -lean_ctor_set(x_207, 0, x_206); -lean_ctor_set(x_207, 1, x_205); -x_208 = l_Lean_Meta_Grind_addCongrTable___closed__7; -if (lean_is_scalar(x_188)) { - x_209 = lean_alloc_ctor(7, 2, 0); + x_242 = x_233; + lean_ctor_set_tag(x_242, 7); +} +lean_ctor_set(x_242, 0, x_241); +lean_ctor_set(x_242, 1, x_240); +x_243 = l_Lean_Meta_Grind_addCongrTable___closed__7; +if (lean_is_scalar(x_221)) { + x_244 = lean_alloc_ctor(7, 2, 0); } else { - x_209 = x_188; - lean_ctor_set_tag(x_209, 7); + x_244 = x_221; + lean_ctor_set_tag(x_244, 7); +} +lean_ctor_set(x_244, 0, x_242); +lean_ctor_set(x_244, 1, x_243); +x_245 = l_Lean_indentExpr(x_220); +x_246 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_246, 0, x_244); +lean_ctor_set(x_246, 1, x_245); +x_247 = l_Lean_Meta_Grind_addCongrTable___closed__9; +x_248 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_248, 0, x_246); +lean_ctor_set(x_248, 1, x_247); +x_249 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_229, x_248, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_239); +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_249, 1); +lean_inc(x_251); +lean_dec(x_249); +x_252 = lean_apply_10(x_234, x_250, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_251); +return x_252; +} +else +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +lean_dec(x_233); +lean_dec(x_221); +lean_dec(x_220); +lean_dec(x_9); +lean_dec(x_8); +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_253 = lean_ctor_get(x_238, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_238, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_238)) { + lean_ctor_release(x_238, 0); + lean_ctor_release(x_238, 1); + x_255 = x_238; +} else { + lean_dec_ref(x_238); + x_255 = lean_box(0); +} +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(1, 2, 0); +} else { + x_256 = x_255; +} +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_254); +return x_256; } -lean_ctor_set(x_209, 0, x_207); -lean_ctor_set(x_209, 1, x_208); -x_210 = l_Lean_indentExpr(x_187); -x_211 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_211, 0, x_209); -lean_ctor_set(x_211, 1, x_210); -x_212 = l_Lean_Meta_Grind_addCongrTable___closed__9; -x_213 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_213, 0, x_211); -lean_ctor_set(x_213, 1, x_212); -x_214 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_196, x_213, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_199); -x_215 = lean_ctor_get(x_214, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_214, 1); -lean_inc(x_216); -lean_dec(x_214); -x_217 = lean_apply_10(x_201, x_215, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_216); -return x_217; } } else { -lean_object* x_218; lean_object* x_219; lean_object* x_220; -lean_dec(x_188); -x_218 = lean_ctor_get(x_192, 1); -lean_inc(x_218); -lean_dec(x_192); -x_219 = lean_box(0); -x_220 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_187, x_219, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_218); +lean_object* x_257; lean_object* x_258; lean_object* x_259; +lean_dec(x_221); +x_257 = lean_ctor_get(x_225, 1); +lean_inc(x_257); +lean_dec(x_225); +x_258 = lean_box(0); +x_259 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_220, x_258, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_257); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_220; +return x_259; } } else { -lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; -lean_dec(x_188); -lean_dec(x_187); +lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +lean_dec(x_221); +lean_dec(x_220); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2682,41 +2694,41 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_221 = lean_ctor_get(x_192, 0); -lean_inc(x_221); -x_222 = lean_ctor_get(x_192, 1); -lean_inc(x_222); -if (lean_is_exclusive(x_192)) { - lean_ctor_release(x_192, 0); - lean_ctor_release(x_192, 1); - x_223 = x_192; +x_260 = lean_ctor_get(x_225, 0); +lean_inc(x_260); +x_261 = lean_ctor_get(x_225, 1); +lean_inc(x_261); +if (lean_is_exclusive(x_225)) { + lean_ctor_release(x_225, 0); + lean_ctor_release(x_225, 1); + x_262 = x_225; } else { - lean_dec_ref(x_192); - x_223 = lean_box(0); + lean_dec_ref(x_225); + x_262 = lean_box(0); } -if (lean_is_scalar(x_223)) { - x_224 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_262)) { + x_263 = lean_alloc_ctor(1, 2, 0); } else { - x_224 = x_223; + x_263 = x_262; } -lean_ctor_set(x_224, 0, x_221); -lean_ctor_set(x_224, 1, x_222); -return x_224; +lean_ctor_set(x_263, 0, x_260); +lean_ctor_set(x_263, 1, x_261); +return x_263; } } else { -lean_object* x_225; lean_object* x_226; -lean_dec(x_190); -lean_dec(x_189); -lean_dec(x_188); -x_225 = lean_box(0); -x_226 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_187, x_225, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_157); +lean_object* x_264; lean_object* x_265; +lean_dec(x_223); +lean_dec(x_222); +lean_dec(x_221); +x_264 = lean_box(0); +x_265 = l_Lean_Meta_Grind_addCongrTable___lambda__2(x_1, x_220, x_264, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_185); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_226; +return x_265; } } } @@ -2766,38 +2778,6 @@ x_9 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___sp return x_9; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addCongrTable___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -3628,7 +3608,7 @@ return x_36; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; x_37 = lean_ctor_get(x_12, 1); x_38 = lean_ctor_get(x_14, 0); x_39 = lean_ctor_get(x_14, 1); @@ -3636,7 +3616,7 @@ x_40 = lean_ctor_get(x_14, 2); x_41 = lean_ctor_get(x_14, 3); x_42 = lean_ctor_get(x_14, 4); x_43 = lean_ctor_get(x_14, 5); -x_44 = lean_ctor_get_uint8(x_14, sizeof(void*)*14); +x_44 = lean_ctor_get_uint8(x_14, sizeof(void*)*19); x_45 = lean_ctor_get(x_14, 6); x_46 = lean_ctor_get(x_14, 7); x_47 = lean_ctor_get(x_14, 8); @@ -3645,6 +3625,16 @@ x_49 = lean_ctor_get(x_14, 10); x_50 = lean_ctor_get(x_14, 11); x_51 = lean_ctor_get(x_14, 12); x_52 = lean_ctor_get(x_14, 13); +x_53 = lean_ctor_get(x_14, 14); +x_54 = lean_ctor_get(x_14, 15); +x_55 = lean_ctor_get(x_14, 16); +x_56 = lean_ctor_get(x_14, 17); +x_57 = lean_ctor_get(x_14, 18); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); lean_inc(x_52); lean_inc(x_51); lean_inc(x_50); @@ -3661,259 +3651,294 @@ lean_inc(x_39); lean_inc(x_38); lean_dec(x_14); lean_inc(x_42); -x_53 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_42, x_11); -if (lean_obj_tag(x_53) == 0) +x_58 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_42, x_11); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_54 = lean_box(0); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_59 = lean_box(0); lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 1, x_54); +lean_ctor_set(x_12, 1, x_59); lean_ctor_set(x_12, 0, x_1); -x_55 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); -x_56 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_56, 0, x_38); -lean_ctor_set(x_56, 1, x_39); -lean_ctor_set(x_56, 2, x_40); -lean_ctor_set(x_56, 3, x_41); -lean_ctor_set(x_56, 4, x_55); -lean_ctor_set(x_56, 5, x_43); -lean_ctor_set(x_56, 6, x_45); -lean_ctor_set(x_56, 7, x_46); -lean_ctor_set(x_56, 8, x_47); -lean_ctor_set(x_56, 9, x_48); -lean_ctor_set(x_56, 10, x_49); -lean_ctor_set(x_56, 11, x_50); -lean_ctor_set(x_56, 12, x_51); -lean_ctor_set(x_56, 13, x_52); -lean_ctor_set_uint8(x_56, sizeof(void*)*14, x_44); -x_57 = lean_st_ref_set(x_2, x_56, x_37); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_59 = x_57; +x_60 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); +x_61 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_61, 0, x_38); +lean_ctor_set(x_61, 1, x_39); +lean_ctor_set(x_61, 2, x_40); +lean_ctor_set(x_61, 3, x_41); +lean_ctor_set(x_61, 4, x_60); +lean_ctor_set(x_61, 5, x_43); +lean_ctor_set(x_61, 6, x_45); +lean_ctor_set(x_61, 7, x_46); +lean_ctor_set(x_61, 8, x_47); +lean_ctor_set(x_61, 9, x_48); +lean_ctor_set(x_61, 10, x_49); +lean_ctor_set(x_61, 11, x_50); +lean_ctor_set(x_61, 12, x_51); +lean_ctor_set(x_61, 13, x_52); +lean_ctor_set(x_61, 14, x_53); +lean_ctor_set(x_61, 15, x_54); +lean_ctor_set(x_61, 16, x_55); +lean_ctor_set(x_61, 17, x_56); +lean_ctor_set(x_61, 18, x_57); +lean_ctor_set_uint8(x_61, sizeof(void*)*19, x_44); +x_62 = lean_st_ref_set(x_2, x_61, x_37); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_64 = x_62; } else { - lean_dec_ref(x_57); - x_59 = lean_box(0); + lean_dec_ref(x_62); + x_64 = lean_box(0); } -x_60 = lean_box(0); -if (lean_is_scalar(x_59)) { - x_61 = lean_alloc_ctor(0, 2, 0); +x_65 = lean_box(0); +if (lean_is_scalar(x_64)) { + x_66 = lean_alloc_ctor(0, 2, 0); } else { - x_61 = x_59; + x_66 = x_64; } -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_58); -return x_61; +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_63); +return x_66; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_62 = lean_ctor_get(x_53, 0); -lean_inc(x_62); -lean_dec(x_53); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_67 = lean_ctor_get(x_58, 0); +lean_inc(x_67); +lean_dec(x_58); lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 1, x_62); +lean_ctor_set(x_12, 1, x_67); lean_ctor_set(x_12, 0, x_1); -x_63 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); -x_64 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_64, 0, x_38); -lean_ctor_set(x_64, 1, x_39); -lean_ctor_set(x_64, 2, x_40); -lean_ctor_set(x_64, 3, x_41); -lean_ctor_set(x_64, 4, x_63); -lean_ctor_set(x_64, 5, x_43); -lean_ctor_set(x_64, 6, x_45); -lean_ctor_set(x_64, 7, x_46); -lean_ctor_set(x_64, 8, x_47); -lean_ctor_set(x_64, 9, x_48); -lean_ctor_set(x_64, 10, x_49); -lean_ctor_set(x_64, 11, x_50); -lean_ctor_set(x_64, 12, x_51); -lean_ctor_set(x_64, 13, x_52); -lean_ctor_set_uint8(x_64, sizeof(void*)*14, x_44); -x_65 = lean_st_ref_set(x_2, x_64, x_37); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_67 = x_65; +x_68 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_42, x_11, x_12); +x_69 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_69, 0, x_38); +lean_ctor_set(x_69, 1, x_39); +lean_ctor_set(x_69, 2, x_40); +lean_ctor_set(x_69, 3, x_41); +lean_ctor_set(x_69, 4, x_68); +lean_ctor_set(x_69, 5, x_43); +lean_ctor_set(x_69, 6, x_45); +lean_ctor_set(x_69, 7, x_46); +lean_ctor_set(x_69, 8, x_47); +lean_ctor_set(x_69, 9, x_48); +lean_ctor_set(x_69, 10, x_49); +lean_ctor_set(x_69, 11, x_50); +lean_ctor_set(x_69, 12, x_51); +lean_ctor_set(x_69, 13, x_52); +lean_ctor_set(x_69, 14, x_53); +lean_ctor_set(x_69, 15, x_54); +lean_ctor_set(x_69, 16, x_55); +lean_ctor_set(x_69, 17, x_56); +lean_ctor_set(x_69, 18, x_57); +lean_ctor_set_uint8(x_69, sizeof(void*)*19, x_44); +x_70 = lean_st_ref_set(x_2, x_69, x_37); +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_72 = x_70; } else { - lean_dec_ref(x_65); - x_67 = lean_box(0); + lean_dec_ref(x_70); + x_72 = lean_box(0); } -x_68 = lean_box(0); -if (lean_is_scalar(x_67)) { - x_69 = lean_alloc_ctor(0, 2, 0); +x_73 = lean_box(0); +if (lean_is_scalar(x_72)) { + x_74 = lean_alloc_ctor(0, 2, 0); } else { - x_69 = x_67; + x_74 = x_72; } -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_66); -return x_69; +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_71); +return x_74; } } } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_70 = lean_ctor_get(x_12, 0); -x_71 = lean_ctor_get(x_12, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_12); -x_72 = lean_ctor_get(x_70, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_70, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_70, 2); -lean_inc(x_74); -x_75 = lean_ctor_get(x_70, 3); -lean_inc(x_75); -x_76 = lean_ctor_get(x_70, 4); +lean_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; uint8_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_75 = lean_ctor_get(x_12, 0); +x_76 = lean_ctor_get(x_12, 1); lean_inc(x_76); -x_77 = lean_ctor_get(x_70, 5); +lean_inc(x_75); +lean_dec(x_12); +x_77 = lean_ctor_get(x_75, 0); lean_inc(x_77); -x_78 = lean_ctor_get_uint8(x_70, sizeof(void*)*14); -x_79 = lean_ctor_get(x_70, 6); +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +x_79 = lean_ctor_get(x_75, 2); lean_inc(x_79); -x_80 = lean_ctor_get(x_70, 7); +x_80 = lean_ctor_get(x_75, 3); lean_inc(x_80); -x_81 = lean_ctor_get(x_70, 8); +x_81 = lean_ctor_get(x_75, 4); lean_inc(x_81); -x_82 = lean_ctor_get(x_70, 9); +x_82 = lean_ctor_get(x_75, 5); lean_inc(x_82); -x_83 = lean_ctor_get(x_70, 10); -lean_inc(x_83); -x_84 = lean_ctor_get(x_70, 11); +x_83 = lean_ctor_get_uint8(x_75, sizeof(void*)*19); +x_84 = lean_ctor_get(x_75, 6); lean_inc(x_84); -x_85 = lean_ctor_get(x_70, 12); +x_85 = lean_ctor_get(x_75, 7); lean_inc(x_85); -x_86 = lean_ctor_get(x_70, 13); +x_86 = lean_ctor_get(x_75, 8); lean_inc(x_86); -if (lean_is_exclusive(x_70)) { - lean_ctor_release(x_70, 0); - lean_ctor_release(x_70, 1); - lean_ctor_release(x_70, 2); - lean_ctor_release(x_70, 3); - lean_ctor_release(x_70, 4); - lean_ctor_release(x_70, 5); - lean_ctor_release(x_70, 6); - lean_ctor_release(x_70, 7); - lean_ctor_release(x_70, 8); - lean_ctor_release(x_70, 9); - lean_ctor_release(x_70, 10); - lean_ctor_release(x_70, 11); - lean_ctor_release(x_70, 12); - lean_ctor_release(x_70, 13); - x_87 = x_70; +x_87 = lean_ctor_get(x_75, 9); +lean_inc(x_87); +x_88 = lean_ctor_get(x_75, 10); +lean_inc(x_88); +x_89 = lean_ctor_get(x_75, 11); +lean_inc(x_89); +x_90 = lean_ctor_get(x_75, 12); +lean_inc(x_90); +x_91 = lean_ctor_get(x_75, 13); +lean_inc(x_91); +x_92 = lean_ctor_get(x_75, 14); +lean_inc(x_92); +x_93 = lean_ctor_get(x_75, 15); +lean_inc(x_93); +x_94 = lean_ctor_get(x_75, 16); +lean_inc(x_94); +x_95 = lean_ctor_get(x_75, 17); +lean_inc(x_95); +x_96 = lean_ctor_get(x_75, 18); +lean_inc(x_96); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + lean_ctor_release(x_75, 2); + lean_ctor_release(x_75, 3); + lean_ctor_release(x_75, 4); + lean_ctor_release(x_75, 5); + lean_ctor_release(x_75, 6); + lean_ctor_release(x_75, 7); + lean_ctor_release(x_75, 8); + lean_ctor_release(x_75, 9); + lean_ctor_release(x_75, 10); + lean_ctor_release(x_75, 11); + lean_ctor_release(x_75, 12); + lean_ctor_release(x_75, 13); + lean_ctor_release(x_75, 14); + lean_ctor_release(x_75, 15); + lean_ctor_release(x_75, 16); + lean_ctor_release(x_75, 17); + lean_ctor_release(x_75, 18); + x_97 = x_75; } else { - lean_dec_ref(x_70); - x_87 = lean_box(0); + lean_dec_ref(x_75); + x_97 = lean_box(0); } -lean_inc(x_76); -x_88 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_76, x_11); -if (lean_obj_tag(x_88) == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_89 = lean_box(0); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_1); -lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_76, x_11, x_90); -if (lean_is_scalar(x_87)) { - x_92 = lean_alloc_ctor(0, 14, 1); +lean_inc(x_81); +x_98 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__1(x_81, x_11); +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_99 = lean_box(0); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_1); +lean_ctor_set(x_100, 1, x_99); +x_101 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_81, x_11, x_100); +if (lean_is_scalar(x_97)) { + x_102 = lean_alloc_ctor(0, 19, 1); } else { - x_92 = x_87; -} -lean_ctor_set(x_92, 0, x_72); -lean_ctor_set(x_92, 1, x_73); -lean_ctor_set(x_92, 2, x_74); -lean_ctor_set(x_92, 3, x_75); -lean_ctor_set(x_92, 4, x_91); -lean_ctor_set(x_92, 5, x_77); -lean_ctor_set(x_92, 6, x_79); -lean_ctor_set(x_92, 7, x_80); -lean_ctor_set(x_92, 8, x_81); -lean_ctor_set(x_92, 9, x_82); -lean_ctor_set(x_92, 10, x_83); -lean_ctor_set(x_92, 11, x_84); -lean_ctor_set(x_92, 12, x_85); -lean_ctor_set(x_92, 13, x_86); -lean_ctor_set_uint8(x_92, sizeof(void*)*14, x_78); -x_93 = lean_st_ref_set(x_2, x_92, x_71); -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_95 = x_93; + x_102 = x_97; +} +lean_ctor_set(x_102, 0, x_77); +lean_ctor_set(x_102, 1, x_78); +lean_ctor_set(x_102, 2, x_79); +lean_ctor_set(x_102, 3, x_80); +lean_ctor_set(x_102, 4, x_101); +lean_ctor_set(x_102, 5, x_82); +lean_ctor_set(x_102, 6, x_84); +lean_ctor_set(x_102, 7, x_85); +lean_ctor_set(x_102, 8, x_86); +lean_ctor_set(x_102, 9, x_87); +lean_ctor_set(x_102, 10, x_88); +lean_ctor_set(x_102, 11, x_89); +lean_ctor_set(x_102, 12, x_90); +lean_ctor_set(x_102, 13, x_91); +lean_ctor_set(x_102, 14, x_92); +lean_ctor_set(x_102, 15, x_93); +lean_ctor_set(x_102, 16, x_94); +lean_ctor_set(x_102, 17, x_95); +lean_ctor_set(x_102, 18, x_96); +lean_ctor_set_uint8(x_102, sizeof(void*)*19, x_83); +x_103 = lean_st_ref_set(x_2, x_102, x_76); +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_105 = x_103; } else { - lean_dec_ref(x_93); - x_95 = lean_box(0); + lean_dec_ref(x_103); + x_105 = lean_box(0); } -x_96 = lean_box(0); -if (lean_is_scalar(x_95)) { - x_97 = lean_alloc_ctor(0, 2, 0); +x_106 = lean_box(0); +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); } else { - x_97 = x_95; + x_107 = x_105; } -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_94); -return x_97; +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_104); +return x_107; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_98 = lean_ctor_get(x_88, 0); -lean_inc(x_98); -lean_dec(x_88); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_1); -lean_ctor_set(x_99, 1, x_98); -x_100 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_76, x_11, x_99); -if (lean_is_scalar(x_87)) { - x_101 = lean_alloc_ctor(0, 14, 1); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_108 = lean_ctor_get(x_98, 0); +lean_inc(x_108); +lean_dec(x_98); +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_1); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap___spec__4(x_81, x_11, x_109); +if (lean_is_scalar(x_97)) { + x_111 = lean_alloc_ctor(0, 19, 1); } else { - x_101 = x_87; -} -lean_ctor_set(x_101, 0, x_72); -lean_ctor_set(x_101, 1, x_73); -lean_ctor_set(x_101, 2, x_74); -lean_ctor_set(x_101, 3, x_75); -lean_ctor_set(x_101, 4, x_100); -lean_ctor_set(x_101, 5, x_77); -lean_ctor_set(x_101, 6, x_79); -lean_ctor_set(x_101, 7, x_80); -lean_ctor_set(x_101, 8, x_81); -lean_ctor_set(x_101, 9, x_82); -lean_ctor_set(x_101, 10, x_83); -lean_ctor_set(x_101, 11, x_84); -lean_ctor_set(x_101, 12, x_85); -lean_ctor_set(x_101, 13, x_86); -lean_ctor_set_uint8(x_101, sizeof(void*)*14, x_78); -x_102 = lean_st_ref_set(x_2, x_101, x_71); -x_103 = lean_ctor_get(x_102, 1); -lean_inc(x_103); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_104 = x_102; + x_111 = x_97; +} +lean_ctor_set(x_111, 0, x_77); +lean_ctor_set(x_111, 1, x_78); +lean_ctor_set(x_111, 2, x_79); +lean_ctor_set(x_111, 3, x_80); +lean_ctor_set(x_111, 4, x_110); +lean_ctor_set(x_111, 5, x_82); +lean_ctor_set(x_111, 6, x_84); +lean_ctor_set(x_111, 7, x_85); +lean_ctor_set(x_111, 8, x_86); +lean_ctor_set(x_111, 9, x_87); +lean_ctor_set(x_111, 10, x_88); +lean_ctor_set(x_111, 11, x_89); +lean_ctor_set(x_111, 12, x_90); +lean_ctor_set(x_111, 13, x_91); +lean_ctor_set(x_111, 14, x_92); +lean_ctor_set(x_111, 15, x_93); +lean_ctor_set(x_111, 16, x_94); +lean_ctor_set(x_111, 17, x_95); +lean_ctor_set(x_111, 18, x_96); +lean_ctor_set_uint8(x_111, sizeof(void*)*19, x_83); +x_112 = lean_st_ref_set(x_2, x_111, x_76); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_114 = x_112; } else { - lean_dec_ref(x_102); - x_104 = lean_box(0); + lean_dec_ref(x_112); + x_114 = lean_box(0); } -x_105 = lean_box(0); -if (lean_is_scalar(x_104)) { - x_106 = lean_alloc_ctor(0, 2, 0); +x_115 = lean_box(0); +if (lean_is_scalar(x_114)) { + x_116 = lean_alloc_ctor(0, 2, 0); } else { - x_106 = x_104; + x_116 = x_114; } -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_103); -return x_106; +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_113); +return x_116; } } } @@ -3989,168 +4014,364 @@ lean_dec(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_14; -x_14 = lean_usize_dec_lt(x_3, x_2); -if (x_14 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_15; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_4); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -else +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_12, 0); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_array_uget(x_4, x_3); -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_array_uset(x_4, x_3, x_17); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_16, x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_19) == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_12, 1); +x_17 = lean_ctor_get(x_14, 16); +lean_ctor_set_tag(x_12, 1); +lean_ctor_set(x_12, 1, x_17); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_14, 16, x_12); +x_18 = lean_st_ref_set(x_3, x_14, x_16); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = 1; -x_23 = lean_usize_add(x_3, x_22); -x_24 = lean_array_uset(x_18, x_3, x_20); -x_3 = x_23; -x_4 = x_24; -x_13 = x_21; -goto _start; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; } else { -uint8_t x_26; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); lean_dec(x_18); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) -{ -return x_19; +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_19, 0); -x_28 = lean_ctor_get(x_19, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_25 = lean_ctor_get(x_12, 1); +x_26 = lean_ctor_get(x_14, 0); +x_27 = lean_ctor_get(x_14, 1); +x_28 = lean_ctor_get(x_14, 2); +x_29 = lean_ctor_get(x_14, 3); +x_30 = lean_ctor_get(x_14, 4); +x_31 = lean_ctor_get(x_14, 5); +x_32 = lean_ctor_get_uint8(x_14, sizeof(void*)*19); +x_33 = lean_ctor_get(x_14, 6); +x_34 = lean_ctor_get(x_14, 7); +x_35 = lean_ctor_get(x_14, 8); +x_36 = lean_ctor_get(x_14, 9); +x_37 = lean_ctor_get(x_14, 10); +x_38 = lean_ctor_get(x_14, 11); +x_39 = lean_ctor_get(x_14, 12); +x_40 = lean_ctor_get(x_14, 13); +x_41 = lean_ctor_get(x_14, 14); +x_42 = lean_ctor_get(x_14, 15); +x_43 = lean_ctor_get(x_14, 16); +x_44 = lean_ctor_get(x_14, 17); +x_45 = lean_ctor_get(x_14, 18); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); -lean_dec(x_19); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_inc(x_26); +lean_dec(x_14); +lean_ctor_set_tag(x_12, 1); +lean_ctor_set(x_12, 1, x_43); +lean_ctor_set(x_12, 0, x_1); +x_46 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_46, 0, x_26); +lean_ctor_set(x_46, 1, x_27); +lean_ctor_set(x_46, 2, x_28); +lean_ctor_set(x_46, 3, x_29); +lean_ctor_set(x_46, 4, x_30); +lean_ctor_set(x_46, 5, x_31); +lean_ctor_set(x_46, 6, x_33); +lean_ctor_set(x_46, 7, x_34); +lean_ctor_set(x_46, 8, x_35); +lean_ctor_set(x_46, 9, x_36); +lean_ctor_set(x_46, 10, x_37); +lean_ctor_set(x_46, 11, x_38); +lean_ctor_set(x_46, 12, x_39); +lean_ctor_set(x_46, 13, x_40); +lean_ctor_set(x_46, 14, x_41); +lean_ctor_set(x_46, 15, x_42); +lean_ctor_set(x_46, 16, x_12); +lean_ctor_set(x_46, 17, x_44); +lean_ctor_set(x_46, 18, x_45); +lean_ctor_set_uint8(x_46, sizeof(void*)*19, x_32); +x_47 = lean_st_ref_set(x_3, x_46, x_25); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); +} +x_50 = lean_box(0); +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_49; +} +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_52 = lean_ctor_get(x_12, 0); +x_53 = lean_ctor_get(x_12, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_12); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_55); +x_56 = lean_ctor_get(x_52, 2); +lean_inc(x_56); +x_57 = lean_ctor_get(x_52, 3); +lean_inc(x_57); +x_58 = lean_ctor_get(x_52, 4); +lean_inc(x_58); +x_59 = lean_ctor_get(x_52, 5); +lean_inc(x_59); +x_60 = lean_ctor_get_uint8(x_52, sizeof(void*)*19); +x_61 = lean_ctor_get(x_52, 6); +lean_inc(x_61); +x_62 = lean_ctor_get(x_52, 7); +lean_inc(x_62); +x_63 = lean_ctor_get(x_52, 8); +lean_inc(x_63); +x_64 = lean_ctor_get(x_52, 9); +lean_inc(x_64); +x_65 = lean_ctor_get(x_52, 10); +lean_inc(x_65); +x_66 = lean_ctor_get(x_52, 11); +lean_inc(x_66); +x_67 = lean_ctor_get(x_52, 12); +lean_inc(x_67); +x_68 = lean_ctor_get(x_52, 13); +lean_inc(x_68); +x_69 = lean_ctor_get(x_52, 14); +lean_inc(x_69); +x_70 = lean_ctor_get(x_52, 15); +lean_inc(x_70); +x_71 = lean_ctor_get(x_52, 16); +lean_inc(x_71); +x_72 = lean_ctor_get(x_52, 17); +lean_inc(x_72); +x_73 = lean_ctor_get(x_52, 18); +lean_inc(x_73); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + lean_ctor_release(x_52, 2); + lean_ctor_release(x_52, 3); + lean_ctor_release(x_52, 4); + lean_ctor_release(x_52, 5); + lean_ctor_release(x_52, 6); + lean_ctor_release(x_52, 7); + lean_ctor_release(x_52, 8); + lean_ctor_release(x_52, 9); + lean_ctor_release(x_52, 10); + lean_ctor_release(x_52, 11); + lean_ctor_release(x_52, 12); + lean_ctor_release(x_52, 13); + lean_ctor_release(x_52, 14); + lean_ctor_release(x_52, 15); + lean_ctor_release(x_52, 16); + lean_ctor_release(x_52, 17); + lean_ctor_release(x_52, 18); + x_74 = x_52; +} else { + lean_dec_ref(x_52); + x_74 = lean_box(0); +} +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_1); +lean_ctor_set(x_75, 1, x_71); +if (lean_is_scalar(x_74)) { + x_76 = lean_alloc_ctor(0, 19, 1); +} else { + x_76 = x_74; } +lean_ctor_set(x_76, 0, x_54); +lean_ctor_set(x_76, 1, x_55); +lean_ctor_set(x_76, 2, x_56); +lean_ctor_set(x_76, 3, x_57); +lean_ctor_set(x_76, 4, x_58); +lean_ctor_set(x_76, 5, x_59); +lean_ctor_set(x_76, 6, x_61); +lean_ctor_set(x_76, 7, x_62); +lean_ctor_set(x_76, 8, x_63); +lean_ctor_set(x_76, 9, x_64); +lean_ctor_set(x_76, 10, x_65); +lean_ctor_set(x_76, 11, x_66); +lean_ctor_set(x_76, 12, x_67); +lean_ctor_set(x_76, 13, x_68); +lean_ctor_set(x_76, 14, x_69); +lean_ctor_set(x_76, 15, x_70); +lean_ctor_set(x_76, 16, x_75); +lean_ctor_set(x_76, 17, x_72); +lean_ctor_set(x_76, 18, x_73); +lean_ctor_set_uint8(x_76, sizeof(void*)*19, x_60); +x_77 = lean_st_ref_set(x_3, x_76, x_53); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_79 = x_77; +} else { + lean_dec_ref(x_77); + x_79 = lean_box(0); +} +x_80 = lean_box(0); +if (lean_is_scalar(x_79)) { + x_81 = lean_alloc_ctor(0, 2, 0); +} else { + x_81 = x_79; } +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_78); +return x_81; } } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1() { _start: { -if (lean_obj_tag(x_2) == 5) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2() { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_2, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_2, 1); -lean_inc(x_15); -lean_dec(x_2); -x_16 = lean_array_set(x_3, x_4, x_15); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_4, x_17); -lean_dec(x_4); -x_2 = x_14; -x_3 = x_16; -x_4 = x_18; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("candidate", 9, 9); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3() { +_start: { -size_t x_20; size_t x_21; lean_object* x_22; -lean_dec(x_4); -x_20 = lean_array_size(x_3); -x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(x_1, x_20, x_21, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_22) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1; +x_3 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3; +x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_22, 0); -x_25 = l_Lean_mkAppN(x_2, x_24); -lean_dec(x_24); -lean_ctor_set(x_22, 0, x_25); -return x_22; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(x_1, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_17; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_22, 0); -x_27 = lean_ctor_get(x_22, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_22); -x_28 = l_Lean_mkAppN(x_2, x_26); +uint8_t x_18; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_12, 1); +x_20 = lean_ctor_get(x_12, 0); +lean_dec(x_20); +x_21 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +lean_inc(x_1); +x_23 = l_Lean_MessageData_ofExpr(x_1); +x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_23); +lean_ctor_set(x_12, 0, x_24); +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_12); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_11, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); lean_dec(x_26); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); +x_29 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(x_1, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); +lean_dec(x_27); return x_29; } -} else { uint8_t x_30; -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_22); +lean_free_object(x_12); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_21); if (x_30 == 0) { -return x_22; +return x_21; } else { lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_22, 0); -x_32 = lean_ctor_get(x_22, 1); +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_22); +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); @@ -4158,1399 +4379,1406 @@ return x_33; } } } +else +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_12, 1); +lean_inc(x_34); +lean_dec(x_12); +x_35 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +lean_inc(x_1); +x_37 = l_Lean_MessageData_ofExpr(x_1); +x_38 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_11, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_36); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(x_1, x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +lean_dec(x_42); +return x_44; } +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_dec(x_1); +x_45 = lean_ctor_get(x_35, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_35, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_47 = x_35; +} else { + lean_dec_ref(x_35); + x_47 = lean_box(0); } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1() { +if (lean_is_scalar(x_47)) { + x_48 = lean_alloc_ctor(1, 2, 0); +} else { + x_48 = x_47; +} +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); +x_1 = lean_mk_string_unchecked("Eq", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); +x_1 = lean_mk_string_unchecked("HEq", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed), 4, 0); +x_1 = lean_mk_string_unchecked("True", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed), 4, 0); +x_1 = lean_mk_string_unchecked("False", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8() { _start: { -uint8_t x_12; -x_12 = l_Lean_Expr_isBVar(x_1); -if (x_12 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9() { +_start: { -lean_object* x_13; uint8_t x_14; -x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; -x_14 = lean_expr_eqv(x_1, x_13); -if (x_14 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10() { +_start: { -lean_object* x_15; -x_15 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); -if (lean_obj_tag(x_15) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11() { +_start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_16); -x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_Meta_isMatcherAppCore(x_14, x_1); +lean_dec(x_14); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); lean_inc(x_17); -x_19 = lean_mk_array(x_17, x_18); -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_sub(x_17, x_20); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); lean_dec(x_17); -x_22 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__2(x_2, x_1, x_19, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_20 = l_Lean_Meta_isMatcherAppCore(x_19, x_1); +lean_dec(x_19); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); return x_22; } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_dec(x_1); -x_23 = lean_ctor_get(x_15, 0); -lean_inc(x_23); -lean_dec(x_15); -x_24 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2; -x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3; +lean_object* x_13; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -x_26 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_24, x_25, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_26) == 0) +x_13 = l_Lean_Meta_isInductivePredicate(x_1, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4; -x_30 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5; -lean_inc(x_10); -lean_inc(x_9); -x_31 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_27, x_29, x_30, x_9, x_10, x_28); -if (lean_obj_tag(x_31) == 0) +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_34 = l_Lean_Meta_Grind_canon(x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_33); -if (lean_obj_tag(x_34) == 0) +uint8_t x_16; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = l_Lean_Meta_Grind_shareCommon(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -lean_inc(x_38); -x_40 = l_Lean_Meta_Grind_internalize(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_39); -if (lean_obj_tag(x_40) == 0) -{ -uint8_t x_41; -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_40, 0); -lean_dec(x_42); -x_43 = l_Lean_Meta_Grind_mkGroundPattern(x_38); -lean_ctor_set(x_40, 0, x_43); -return x_40; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_40, 1); -lean_inc(x_44); -lean_dec(x_40); -x_45 = l_Lean_Meta_Grind_mkGroundPattern(x_38); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_44); -return x_46; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); +x_23 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_23; } } else { -uint8_t x_47; -lean_dec(x_38); -x_47 = !lean_is_exclusive(x_40); -if (x_47 == 0) +uint8_t x_24; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +x_24 = !lean_is_exclusive(x_13); +if (x_24 == 0) { -return x_40; +return x_13; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_40, 0); -x_49 = lean_ctor_get(x_40, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_40); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} } } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = l_Lean_Expr_isIte(x_1); +if (x_12 == 0) +{ +uint8_t x_13; +x_13 = l_Lean_Expr_isDIte(x_1); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_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; +x_17 = !lean_is_exclusive(x_14); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_14, 1); +x_19 = lean_ctor_get(x_14, 0); +lean_dec(x_19); +x_20 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_20) == 4) +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes; +x_23 = l_List_elem___at_Lean_addAliasEntry___spec__16(x_21, x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +lean_free_object(x_14); +x_24 = lean_box(0); +x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1(x_21, x_1, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +lean_dec(x_21); +return x_25; +} else { -uint8_t x_51; +lean_object* x_26; +lean_dec(x_21); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_51 = !lean_is_exclusive(x_34); -if (x_51 == 0) -{ -return x_34; +lean_dec(x_1); +x_26 = lean_box(0); +lean_ctor_set(x_14, 0, x_26); +return x_14; +} } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_34, 0); -x_53 = lean_ctor_get(x_34, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_34); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +lean_object* x_27; +lean_dec(x_20); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_1); +x_27 = lean_box(0); +lean_ctor_set(x_14, 0, x_27); +return x_14; } } +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_14, 1); +lean_inc(x_28); +lean_dec(x_14); +x_29 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_29) == 4) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes; +x_32 = l_List_elem___at_Lean_addAliasEntry___spec__16(x_30, x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_box(0); +x_34 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1(x_30, x_1, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); +lean_dec(x_30); +return x_34; } else { -uint8_t x_55; +lean_object* x_35; lean_object* x_36; +lean_dec(x_30); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_55 = !lean_is_exclusive(x_31); -if (x_55 == 0) -{ -return x_31; +lean_dec(x_1); +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_28); +return x_36; +} } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_31, 0); -x_57 = lean_ctor_get(x_31, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_31); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; +lean_object* x_37; lean_object* x_38; +lean_dec(x_29); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_1); +x_37 = lean_box(0); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_28); +return x_38; } } } else { -uint8_t x_59; +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_14, 1); +lean_inc(x_39); +lean_dec(x_14); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_1); +x_40 = l_Lean_Meta_reduceMatcher_x3f(x_1, x_7, x_8, x_9, x_10, x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +if (lean_obj_tag(x_41) == 0) +{ +uint8_t x_42; +lean_dec(x_41); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_59 = !lean_is_exclusive(x_26); -if (x_59 == 0) +lean_dec(x_1); +x_42 = !lean_is_exclusive(x_40); +if (x_42 == 0) { -return x_26; +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 0); +lean_dec(x_43); +x_44 = lean_box(0); +lean_ctor_set(x_40, 0, x_44); +return x_40; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_26, 0); -x_61 = lean_ctor_get(x_26, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_26); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; -} -} +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_box(0); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; } } else { -lean_object* x_63; +lean_object* x_48; lean_object* x_49; +lean_dec(x_41); +x_48 = lean_ctor_get(x_40, 1); +lean_inc(x_48); +lean_dec(x_40); +x_49 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_48); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_1); -lean_ctor_set(x_63, 1, x_11); -return x_63; +return x_49; } } else { -lean_object* x_64; +uint8_t x_50; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_1); +x_50 = !lean_is_exclusive(x_40); +if (x_50 == 0) +{ +return x_40; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_40, 0); +x_52 = lean_ctor_get(x_40, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_40); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +} +else +{ +lean_object* x_54; +x_54 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_54; +} +} +else +{ +lean_object* x_55; +x_55 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = l_Lean_Expr_isApp(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_box(0); +x_15 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2(x_1, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_1); -lean_ctor_set(x_64, 1, x_11); -return x_64; -} +lean_dec(x_1); +return x_11; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; } } -static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; -x_2 = l_instInhabitedPUnit; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3() { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; -x_12 = lean_panic_fn(x_11, x_1); -x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eqRecOn_heq", 11, 11); +return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4() { _start: { -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_5, 1); -x_20 = lean_nat_dec_lt(x_7, x_19); -if (x_20 == 0) -{ -lean_object* x_21; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_6); -lean_ctor_set(x_21, 1, x_18); -return x_21; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: { -lean_object* x_22; lean_object* x_23; -lean_dec(x_6); -x_22 = lean_array_fget(x_3, x_7); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_2); -lean_inc(x_22); -x_23 = l_Lean_Meta_Grind_internalize(x_22, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_23) == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_18 = l_Lean_Expr_constLevels_x21(x_2); +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4; +x_20 = l_Lean_Expr_const___override(x_19, x_18); +lean_inc(x_8); +x_21 = l_Lean_mkApp6(x_20, x_3, x_4, x_5, x_6, x_7, x_8); +x_22 = 1; +x_23 = l_Lean_Meta_Grind_pushEqCore(x_1, x_8, x_21, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_23; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1() { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -lean_inc(x_1); -x_25 = l_Lean_Meta_Grind_registerParent(x_1, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_24); -lean_dec(x_22); -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_ctor_get(x_5, 2); -x_28 = lean_nat_add(x_7, x_27); -lean_dec(x_7); -x_29 = lean_box(0); -x_6 = x_29; -x_7 = x_28; -x_8 = lean_box(0); -x_9 = lean_box(0); -x_18 = x_26; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eqNDRec_heq", 11, 11); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2() { +_start: { -uint8_t x_31; -lean_dec(x_22); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_23); -if (x_31 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: { +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_18 = l_Lean_Expr_constLevels_x21(x_2); +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2; +x_20 = l_Lean_Expr_const___override(x_19, x_18); +lean_inc(x_6); +x_21 = l_Lean_mkApp6(x_20, x_3, x_4, x_5, x_6, x_7, x_8); +x_22 = 1; +x_23 = l_Lean_Meta_Grind_pushEqCore(x_1, x_6, x_21, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); return x_23; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_23, 0); -x_33 = lean_ctor_get(x_23, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_23); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eqRec_heq", 9, 9); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_18 = l_Lean_Expr_constLevels_x21(x_2); +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2; +x_20 = l_Lean_Expr_const___override(x_19, x_18); +lean_inc(x_6); +x_21 = l_Lean_mkApp6(x_20, x_3, x_4, x_5, x_6, x_7, x_8); +x_22 = 1; +x_23 = l_Lean_Meta_Grind_pushEqCore(x_1, x_6, x_21, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_23; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1() { _start: { -lean_object* x_13; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_13 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_15 = l_Lean_Meta_Grind_addCongrTable(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -lean_inc(x_1); -x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -return x_19; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cast_heq", 8, 8); +return x_1; } -else -{ -uint8_t x_20; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) -{ -return x_15; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2() { +_start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else -{ -uint8_t x_24; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_13); -if (x_24 == 0) -{ -return x_13; } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_13, 0); -x_26 = lean_ctor_get(x_13, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_13); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_16 = l_Lean_Expr_constLevels_x21(x_2); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2; +x_18 = l_Lean_Expr_const___override(x_17, x_16); +lean_inc(x_6); +x_19 = l_Lean_mkApp4(x_18, x_3, x_4, x_5, x_6); +x_20 = 1; +x_21 = l_Lean_Meta_Grind_pushEqCore(x_1, x_6, x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_21; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cast", 4, 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2() { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_15 = lean_array_get_size(x_1); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_15); -lean_ctor_set(x_18, 2, x_17); -x_19 = lean_box(0); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_20 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(x_2, x_3, x_1, x_18, x_18, x_19, x_16, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_18); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_apply_10(x_4, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); -return x_22; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else -{ -uint8_t x_23; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_23 = !lean_is_exclusive(x_20); -if (x_23 == 0) -{ -return x_20; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3() { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 0); -x_25 = lean_ctor_get(x_20, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_20); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("recOn", 5, 5); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +x_1 = lean_mk_string_unchecked("ndrec", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); +x_1 = lean_mk_string_unchecked("rec", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; -x_2 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; -x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -switch (lean_obj_tag(x_3)) { -case 0: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_28; uint8_t x_29; -lean_dec(x_5); -lean_inc(x_2); +lean_object* x_11; uint8_t x_12; lean_inc(x_1); -x_15 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); -x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_29 = l_Lean_Expr_isConstOf(x_3, x_28); -if (x_29 == 0) +x_11 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_6, x_7, x_8, x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_30; -x_30 = lean_box(0); -x_16 = x_30; -goto block_27; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = l_Lean_Expr_cleanupAnnotations(x_13); +x_16 = l_Lean_Expr_isApp(x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_15); +lean_dec(x_1); +x_17 = lean_box(0); +lean_ctor_set(x_11, 0, x_17); +return x_11; } else { -lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_31 = lean_array_get_size(x_4); -x_32 = lean_unsigned_to_nat(2u); -x_33 = lean_nat_dec_eq(x_31, x_32); -if (x_33 == 0) +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Lean_Expr_appArg(x_15, lean_box(0)); +x_19 = l_Lean_Expr_appFnCleanup(x_15, lean_box(0)); +x_20 = l_Lean_Expr_isApp(x_19); +if (x_20 == 0) { -lean_object* x_34; -lean_dec(x_31); -x_34 = lean_box(0); -x_16 = x_34; -goto block_27; +lean_object* x_21; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_1); +x_21 = lean_box(0); +lean_ctor_set(x_11, 0, x_21); +return x_11; } else { -lean_object* x_35; uint8_t x_36; -lean_dec(x_15); -lean_dec(x_3); -x_35 = lean_unsigned_to_nat(0u); -x_36 = lean_nat_dec_lt(x_35, x_31); -lean_dec(x_31); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_4); -x_37 = l_Lean_instInhabitedExpr; -x_38 = l_outOfBounds___rarg(x_37); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_38); -x_39 = l_Lean_Meta_Grind_internalize(x_38, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_39) == 0) +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = l_Lean_Expr_appArg(x_19, lean_box(0)); +x_23 = l_Lean_Expr_appFnCleanup(x_19, lean_box(0)); +x_24 = l_Lean_Expr_isApp(x_23); +if (x_24 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -lean_inc(x_1); -x_41 = l_Lean_Meta_Grind_registerParent(x_1, x_38, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_40); -lean_dec(x_38); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_42, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_43); -lean_dec(x_42); -return x_44; +lean_object* x_25; +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_1); +x_25 = lean_box(0); +lean_ctor_set(x_11, 0, x_25); +return x_11; } else { -uint8_t x_45; -lean_dec(x_38); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_45 = !lean_is_exclusive(x_39); -if (x_45 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Expr_appArg(x_23, lean_box(0)); +x_27 = l_Lean_Expr_appFnCleanup(x_23, lean_box(0)); +x_28 = l_Lean_Expr_isApp(x_27); +if (x_28 == 0) { -return x_39; +lean_object* x_29; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_1); +x_29 = lean_box(0); +lean_ctor_set(x_11, 0, x_29); +return x_11; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_39, 0); -x_47 = lean_ctor_get(x_39, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_39); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; -} -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = l_Lean_Expr_appArg(x_27, lean_box(0)); +x_31 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); +x_32 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2; +x_33 = l_Lean_Expr_isConstOf(x_31, x_32); +if (x_33 == 0) +{ +uint8_t x_34; +x_34 = l_Lean_Expr_isApp(x_31); +if (x_34 == 0) +{ +lean_object* x_35; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_1); +x_35 = lean_box(0); +lean_ctor_set(x_11, 0, x_35); +return x_11; } else { -lean_object* x_49; lean_object* x_50; -x_49 = lean_array_fget(x_4, x_35); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_49); -x_50 = l_Lean_Meta_Grind_internalize(x_49, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_50) == 0) +lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_36 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_37 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_38 = l_Lean_Expr_isApp(x_37); +if (x_38 == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -lean_dec(x_50); -lean_inc(x_1); -x_52 = l_Lean_Meta_Grind_registerParent(x_1, x_49, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_51); -lean_dec(x_49); -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); -lean_dec(x_52); -x_55 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_53, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_54); -lean_dec(x_53); -return x_55; +lean_object* x_39; +lean_dec(x_37); +lean_dec(x_36); +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_18); +lean_dec(x_1); +x_39 = lean_box(0); +lean_ctor_set(x_11, 0, x_39); +return x_11; } else { -uint8_t x_56; -lean_dec(x_49); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_40 = l_Lean_Expr_appArg(x_37, lean_box(0)); +x_41 = l_Lean_Expr_appFnCleanup(x_37, lean_box(0)); +x_42 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4; +x_43 = l_Lean_Expr_isConstOf(x_41, x_42); +if (x_43 == 0) +{ +lean_object* x_44; uint8_t x_45; +x_44 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6; +x_45 = l_Lean_Expr_isConstOf(x_41, x_44); +if (x_45 == 0) +{ +lean_object* x_46; uint8_t x_47; +x_46 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8; +x_47 = l_Lean_Expr_isConstOf(x_41, x_46); +if (x_47 == 0) +{ +lean_object* x_48; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_36); +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_18); lean_dec(x_1); -x_56 = !lean_is_exclusive(x_50); -if (x_56 == 0) +x_48 = lean_box(0); +lean_ctor_set(x_11, 0, x_48); +return x_11; +} +else { -return x_50; +lean_object* x_49; +lean_free_object(x_11); +x_49 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3(x_1, x_41, x_40, x_36, x_30, x_26, x_22, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_41); +return x_49; +} } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_50, 0); -x_58 = lean_ctor_get(x_50, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_50); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +lean_object* x_50; +lean_free_object(x_11); +x_50 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(x_1, x_41, x_40, x_36, x_30, x_26, x_22, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_41); +return x_50; } } +else +{ +lean_object* x_51; +lean_free_object(x_11); +x_51 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(x_1, x_41, x_40, x_36, x_30, x_26, x_22, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_41); +return x_51; } } } -block_27: -{ -lean_object* x_17; -lean_dec(x_16); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_17 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -lean_inc(x_1); -x_19 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_18); -lean_dec(x_3); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_15, x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); -lean_dec(x_20); -lean_dec(x_4); -return x_22; } else { -uint8_t x_23; -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_23 = !lean_is_exclusive(x_17); -if (x_23 == 0) -{ -return x_17; +lean_object* x_52; +lean_free_object(x_11); +x_52 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4(x_1, x_31, x_30, x_26, x_22, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_31); +return x_52; } -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_17, 0); -x_25 = lean_ctor_get(x_17, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_17); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; } } } } -case 1: -{ -lean_object* x_60; lean_object* x_61; lean_object* x_73; uint8_t x_74; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_60 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_60, 0, x_1); -lean_closure_set(x_60, 1, x_2); -x_73 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_74 = l_Lean_Expr_isConstOf(x_3, x_73); -if (x_74 == 0) -{ -lean_object* x_75; -x_75 = lean_box(0); -x_61 = x_75; -goto block_72; } else { -lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_76 = lean_array_get_size(x_4); -x_77 = lean_unsigned_to_nat(2u); -x_78 = lean_nat_dec_eq(x_76, x_77); -if (x_78 == 0) +lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_53 = lean_ctor_get(x_11, 0); +x_54 = lean_ctor_get(x_11, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_11); +x_55 = l_Lean_Expr_cleanupAnnotations(x_53); +x_56 = l_Lean_Expr_isApp(x_55); +if (x_56 == 0) { -lean_object* x_79; -lean_dec(x_76); -x_79 = lean_box(0); -x_61 = x_79; -goto block_72; +lean_object* x_57; lean_object* x_58; +lean_dec(x_55); +lean_dec(x_1); +x_57 = lean_box(0); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_54); +return x_58; } else { -lean_object* x_80; uint8_t x_81; -lean_dec(x_60); -lean_dec(x_3); -x_80 = lean_unsigned_to_nat(0u); -x_81 = lean_nat_dec_lt(x_80, x_76); -lean_dec(x_76); -if (x_81 == 0) -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; -lean_dec(x_4); -x_82 = l_Lean_instInhabitedExpr; -x_83 = l_outOfBounds___rarg(x_82); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_83); -x_84 = l_Lean_Meta_Grind_internalize(x_83, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_84) == 0) +lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_59 = l_Lean_Expr_appArg(x_55, lean_box(0)); +x_60 = l_Lean_Expr_appFnCleanup(x_55, lean_box(0)); +x_61 = l_Lean_Expr_isApp(x_60); +if (x_61 == 0) { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_85 = lean_ctor_get(x_84, 1); -lean_inc(x_85); -lean_dec(x_84); -lean_inc(x_1); -x_86 = l_Lean_Meta_Grind_registerParent(x_1, x_83, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_85); -lean_dec(x_83); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); -lean_inc(x_88); -lean_dec(x_86); -x_89 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_87, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_88); -lean_dec(x_87); -return x_89; +lean_object* x_62; lean_object* x_63; +lean_dec(x_60); +lean_dec(x_59); +lean_dec(x_1); +x_62 = lean_box(0); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_54); +return x_63; } else { -uint8_t x_90; -lean_dec(x_83); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_90 = !lean_is_exclusive(x_84); -if (x_90 == 0) +lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_64 = l_Lean_Expr_appArg(x_60, lean_box(0)); +x_65 = l_Lean_Expr_appFnCleanup(x_60, lean_box(0)); +x_66 = l_Lean_Expr_isApp(x_65); +if (x_66 == 0) { -return x_84; +lean_object* x_67; lean_object* x_68; +lean_dec(x_65); +lean_dec(x_64); +lean_dec(x_59); +lean_dec(x_1); +x_67 = lean_box(0); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_54); +return x_68; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_84, 0); -x_92 = lean_ctor_get(x_84, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_84); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; -} -} +lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_69 = l_Lean_Expr_appArg(x_65, lean_box(0)); +x_70 = l_Lean_Expr_appFnCleanup(x_65, lean_box(0)); +x_71 = l_Lean_Expr_isApp(x_70); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_64); +lean_dec(x_59); +lean_dec(x_1); +x_72 = lean_box(0); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_54); +return x_73; } else { -lean_object* x_94; lean_object* x_95; -x_94 = lean_array_fget(x_4, x_80); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_94); -x_95 = l_Lean_Meta_Grind_internalize(x_94, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_95) == 0) +lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_74 = l_Lean_Expr_appArg(x_70, lean_box(0)); +x_75 = l_Lean_Expr_appFnCleanup(x_70, lean_box(0)); +x_76 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2; +x_77 = l_Lean_Expr_isConstOf(x_75, x_76); +if (x_77 == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_96 = lean_ctor_get(x_95, 1); -lean_inc(x_96); -lean_dec(x_95); -lean_inc(x_1); -x_97 = l_Lean_Meta_Grind_registerParent(x_1, x_94, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_96); -lean_dec(x_94); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -x_100 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_98, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_99); -lean_dec(x_98); -return x_100; +uint8_t x_78; +x_78 = l_Lean_Expr_isApp(x_75); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; +lean_dec(x_75); +lean_dec(x_74); +lean_dec(x_69); +lean_dec(x_64); +lean_dec(x_59); +lean_dec(x_1); +x_79 = lean_box(0); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_54); +return x_80; } else { -uint8_t x_101; -lean_dec(x_94); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_101 = !lean_is_exclusive(x_95); -if (x_101 == 0) +lean_object* x_81; lean_object* x_82; uint8_t x_83; +x_81 = l_Lean_Expr_appArg(x_75, lean_box(0)); +x_82 = l_Lean_Expr_appFnCleanup(x_75, lean_box(0)); +x_83 = l_Lean_Expr_isApp(x_82); +if (x_83 == 0) { -return x_95; +lean_object* x_84; lean_object* x_85; +lean_dec(x_82); +lean_dec(x_81); +lean_dec(x_74); +lean_dec(x_69); +lean_dec(x_64); +lean_dec(x_59); +lean_dec(x_1); +x_84 = lean_box(0); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_54); +return x_85; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_95, 0); -x_103 = lean_ctor_get(x_95, 1); -lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_95); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -return x_104; -} -} -} -} -} -block_72: +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = l_Lean_Expr_appArg(x_82, lean_box(0)); +x_87 = l_Lean_Expr_appFnCleanup(x_82, lean_box(0)); +x_88 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4; +x_89 = l_Lean_Expr_isConstOf(x_87, x_88); +if (x_89 == 0) { -lean_object* x_62; -lean_dec(x_61); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_62 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_62) == 0) +lean_object* x_90; uint8_t x_91; +x_90 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6; +x_91 = l_Lean_Expr_isConstOf(x_87, x_90); +if (x_91 == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -lean_dec(x_62); -lean_inc(x_1); -x_64 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_63); -lean_dec(x_3); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); +lean_object* x_92; uint8_t x_93; +x_92 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8; +x_93 = l_Lean_Expr_isConstOf(x_87, x_92); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; +lean_dec(x_87); +lean_dec(x_86); +lean_dec(x_81); +lean_dec(x_74); +lean_dec(x_69); lean_dec(x_64); -x_67 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_60, x_65, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_66); -lean_dec(x_65); -lean_dec(x_4); -return x_67; +lean_dec(x_59); +lean_dec(x_1); +x_94 = lean_box(0); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_54); +return x_95; } else { -uint8_t x_68; -lean_dec(x_60); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_68 = !lean_is_exclusive(x_62); -if (x_68 == 0) -{ -return x_62; +lean_object* x_96; +x_96 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3(x_1, x_87, x_86, x_81, x_74, x_69, x_64, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_54); +lean_dec(x_87); +return x_96; +} } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_62, 0); -x_70 = lean_ctor_get(x_62, 1); -lean_inc(x_70); -lean_inc(x_69); -lean_dec(x_62); -x_71 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -return x_71; +lean_object* x_97; +x_97 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(x_1, x_87, x_86, x_81, x_74, x_69, x_64, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_54); +lean_dec(x_87); +return x_97; } } +else +{ +lean_object* x_98; +x_98 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(x_1, x_87, x_86, x_81, x_74, x_69, x_64, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_54); +lean_dec(x_87); +return x_98; } } -case 2: -{ -lean_object* x_105; lean_object* x_106; lean_object* x_118; uint8_t x_119; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_105 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_105, 0, x_1); -lean_closure_set(x_105, 1, x_2); -x_118 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_119 = l_Lean_Expr_isConstOf(x_3, x_118); -if (x_119 == 0) -{ -lean_object* x_120; -x_120 = lean_box(0); -x_106 = x_120; -goto block_117; } -else -{ -lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_121 = lean_array_get_size(x_4); -x_122 = lean_unsigned_to_nat(2u); -x_123 = lean_nat_dec_eq(x_121, x_122); -if (x_123 == 0) -{ -lean_object* x_124; -lean_dec(x_121); -x_124 = lean_box(0); -x_106 = x_124; -goto block_117; } else { -lean_object* x_125; uint8_t x_126; -lean_dec(x_105); -lean_dec(x_3); -x_125 = lean_unsigned_to_nat(0u); -x_126 = lean_nat_dec_lt(x_125, x_121); -lean_dec(x_121); -if (x_126 == 0) -{ -lean_object* x_127; lean_object* x_128; lean_object* x_129; -lean_dec(x_4); -x_127 = l_Lean_instInhabitedExpr; -x_128 = l_outOfBounds___rarg(x_127); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_128); -x_129 = l_Lean_Meta_Grind_internalize(x_128, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_129) == 0) -{ -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_130 = lean_ctor_get(x_129, 1); -lean_inc(x_130); -lean_dec(x_129); -lean_inc(x_1); -x_131 = l_Lean_Meta_Grind_registerParent(x_1, x_128, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_130); -lean_dec(x_128); -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); -x_134 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_132, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_133); -lean_dec(x_132); -return x_134; +lean_object* x_99; +x_99 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4(x_1, x_75, x_74, x_69, x_64, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_54); +lean_dec(x_75); +return x_99; } -else +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: { -uint8_t x_135; -lean_dec(x_128); +lean_object* x_18; +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); -lean_dec(x_1); -x_135 = !lean_is_exclusive(x_129); -if (x_135 == 0) -{ -return x_129; -} -else -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_136 = lean_ctor_get(x_129, 0); -x_137 = lean_ctor_get(x_129, 1); -lean_inc(x_137); -lean_inc(x_136); -lean_dec(x_129); -x_138 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_138, 0, x_136); -lean_ctor_set(x_138, 1, x_137); -return x_138; -} -} +return x_18; } -else -{ -lean_object* x_139; lean_object* x_140; -x_139 = lean_array_fget(x_4, x_125); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_139); -x_140 = l_Lean_Meta_Grind_internalize(x_139, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_140) == 0) -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_141 = lean_ctor_get(x_140, 1); -lean_inc(x_141); -lean_dec(x_140); -lean_inc(x_1); -x_142 = l_Lean_Meta_Grind_registerParent(x_1, x_139, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_141); -lean_dec(x_139); -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_142, 1); -lean_inc(x_144); -lean_dec(x_142); -x_145 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_143, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_144); -lean_dec(x_143); -return x_145; } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: { -uint8_t x_146; -lean_dec(x_139); +lean_object* x_18; +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); -lean_dec(x_1); -x_146 = !lean_is_exclusive(x_140); -if (x_146 == 0) -{ -return x_140; -} -else -{ -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_140, 0); -x_148 = lean_ctor_get(x_140, 1); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_140); -x_149 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -return x_149; -} -} +return x_18; } } -} -block_117: -{ -lean_object* x_107; -lean_dec(x_106); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_107 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_107) == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_108 = lean_ctor_get(x_107, 1); -lean_inc(x_108); -lean_dec(x_107); -lean_inc(x_1); -x_109 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_108); -lean_dec(x_3); -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_109, 1); -lean_inc(x_111); -lean_dec(x_109); -x_112 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_105, x_110, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_111); -lean_dec(x_110); -lean_dec(x_4); -return x_112; +lean_object* x_18; +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +return x_18; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: { -uint8_t x_113; -lean_dec(x_105); +lean_object* x_16; +x_16 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -5558,112 +5786,34 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_113 = !lean_is_exclusive(x_107); -if (x_113 == 0) -{ -return x_107; -} -else -{ -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_107, 0); -x_115 = lean_ctor_get(x_107, 1); -lean_inc(x_115); -lean_inc(x_114); -lean_dec(x_107); -x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -return x_116; -} -} +return x_16; } } -case 3: +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_150; lean_object* x_151; lean_object* x_163; uint8_t x_164; +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_150 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_150, 0, x_1); -lean_closure_set(x_150, 1, x_2); -x_163 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_164 = l_Lean_Expr_isConstOf(x_3, x_163); -if (x_164 == 0) -{ -lean_object* x_165; -x_165 = lean_box(0); -x_151 = x_165; -goto block_162; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; } -else -{ -lean_object* x_166; lean_object* x_167; uint8_t x_168; -x_166 = lean_array_get_size(x_4); -x_167 = lean_unsigned_to_nat(2u); -x_168 = lean_nat_dec_eq(x_166, x_167); -if (x_168 == 0) -{ -lean_object* x_169; -lean_dec(x_166); -x_169 = lean_box(0); -x_151 = x_169; -goto block_162; } -else -{ -lean_object* x_170; uint8_t x_171; -lean_dec(x_150); -lean_dec(x_3); -x_170 = lean_unsigned_to_nat(0u); -x_171 = lean_nat_dec_lt(x_170, x_166); -lean_dec(x_166); -if (x_171 == 0) -{ -lean_object* x_172; lean_object* x_173; lean_object* x_174; -lean_dec(x_4); -x_172 = l_Lean_instInhabitedExpr; -x_173 = l_outOfBounds___rarg(x_172); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_173); -x_174 = l_Lean_Meta_Grind_internalize(x_173, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_174) == 0) +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_175 = lean_ctor_get(x_174, 1); -lean_inc(x_175); -lean_dec(x_174); -lean_inc(x_1); -x_176 = l_Lean_Meta_Grind_registerParent(x_1, x_173, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_175); -lean_dec(x_173); -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); -lean_inc(x_178); -lean_dec(x_176); -x_179 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_177, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_178); -lean_dec(x_177); -return x_179; -} -else +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_3, x_2); +if (x_14 == 0) { -uint8_t x_180; -lean_dec(x_173); -lean_dec(x_13); +lean_object* x_15; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -5671,34 +5821,19 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); +lean_dec(x_5); lean_dec(x_1); -x_180 = !lean_is_exclusive(x_174); -if (x_180 == 0) -{ -return x_174; -} -else -{ -lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_181 = lean_ctor_get(x_174, 0); -x_182 = lean_ctor_get(x_174, 1); -lean_inc(x_182); -lean_inc(x_181); -lean_dec(x_174); -x_183 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_183, 0, x_181); -lean_ctor_set(x_183, 1, x_182); -return x_183; -} -} +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_4); +lean_ctor_set(x_15, 1, x_13); +return x_15; } else { -lean_object* x_184; lean_object* x_185; -x_184 = lean_array_fget(x_4, x_170); -lean_dec(x_4); -lean_inc(x_13); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_array_uget(x_4, x_3); +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_array_uset(x_4, x_3, x_17); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); @@ -5706,32 +5841,29 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_184); -x_185 = l_Lean_Meta_Grind_internalize(x_184, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_185) == 0) -{ -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_186 = lean_ctor_get(x_185, 1); -lean_inc(x_186); -lean_dec(x_185); +lean_inc(x_5); lean_inc(x_1); -x_187 = l_Lean_Meta_Grind_registerParent(x_1, x_184, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_186); -lean_dec(x_184); -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_187, 1); -lean_inc(x_189); -lean_dec(x_187); -x_190 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_188, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_189); -lean_dec(x_188); -return x_190; +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_16, x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = 1; +x_23 = lean_usize_add(x_3, x_22); +x_24 = lean_array_uset(x_18, x_3, x_20); +x_3 = x_23; +x_4 = x_24; +x_13 = x_21; +goto _start; } else { -uint8_t x_191; -lean_dec(x_184); -lean_dec(x_13); +uint8_t x_26; +lean_dec(x_18); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -5739,1228 +5871,695 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); +lean_dec(x_5); lean_dec(x_1); -x_191 = !lean_is_exclusive(x_185); -if (x_191 == 0) +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) { -return x_185; +return x_19; } else { -lean_object* x_192; lean_object* x_193; lean_object* x_194; -x_192 = lean_ctor_get(x_185, 0); -x_193 = lean_ctor_get(x_185, 1); -lean_inc(x_193); -lean_inc(x_192); -lean_dec(x_185); -x_194 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -return x_194; +lean_object* x_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; } } } } } -block_162: +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -lean_object* x_152; -lean_dec(x_151); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_152 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_152) == 0) +if (lean_obj_tag(x_2) == 5) { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_153 = lean_ctor_get(x_152, 1); -lean_inc(x_153); -lean_dec(x_152); -lean_inc(x_1); -x_154 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_153); -lean_dec(x_3); -x_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); -lean_inc(x_156); -lean_dec(x_154); -x_157 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_150, x_155, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_156); -lean_dec(x_155); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_15); +lean_dec(x_2); +x_16 = lean_array_set(x_3, x_4, x_15); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_4, x_17); lean_dec(x_4); -return x_157; +x_2 = x_14; +x_3 = x_16; +x_4 = x_18; +goto _start; } else { -uint8_t x_158; -lean_dec(x_150); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +size_t x_20; size_t x_21; lean_object* x_22; lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_158 = !lean_is_exclusive(x_152); -if (x_158 == 0) -{ -return x_152; -} -else -{ -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_152, 0); -x_160 = lean_ctor_get(x_152, 1); -lean_inc(x_160); -lean_inc(x_159); -lean_dec(x_152); -x_161 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_161, 0, x_159); -lean_ctor_set(x_161, 1, x_160); -return x_161; -} -} -} -} -case 4: +x_20 = lean_array_size(x_3); +x_21 = 0; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__1(x_1, x_20, x_21, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_207; uint8_t x_208; -lean_dec(x_5); -x_195 = lean_ctor_get(x_3, 0); -lean_inc(x_195); -lean_inc(x_2); -lean_inc(x_1); -x_196 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_196, 0, x_1); -lean_closure_set(x_196, 1, x_2); -x_207 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_208 = l_Lean_Expr_isConstOf(x_3, x_207); -lean_dec(x_3); -if (x_208 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_209; -x_209 = lean_box(0); -x_197 = x_209; -goto block_206; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +x_25 = l_Lean_mkAppN(x_2, x_24); +lean_dec(x_24); +lean_ctor_set(x_22, 0, x_25); +return x_22; } else { -lean_object* x_210; lean_object* x_211; uint8_t x_212; -x_210 = lean_array_get_size(x_4); -x_211 = lean_unsigned_to_nat(2u); -x_212 = lean_nat_dec_eq(x_210, x_211); -if (x_212 == 0) -{ -lean_object* x_213; -lean_dec(x_210); -x_213 = lean_box(0); -x_197 = x_213; -goto block_206; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_22, 0); +x_27 = lean_ctor_get(x_22, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_22); +x_28 = l_Lean_mkAppN(x_2, x_26); +lean_dec(x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } -else -{ -lean_object* x_214; uint8_t x_215; -lean_dec(x_196); -lean_dec(x_195); -x_214 = lean_unsigned_to_nat(0u); -x_215 = lean_nat_dec_lt(x_214, x_210); -lean_dec(x_210); -if (x_215 == 0) -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; -lean_dec(x_4); -x_216 = l_Lean_instInhabitedExpr; -x_217 = l_outOfBounds___rarg(x_216); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_217); -x_218 = l_Lean_Meta_Grind_internalize(x_217, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_218) == 0) -{ -lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_219 = lean_ctor_get(x_218, 1); -lean_inc(x_219); -lean_dec(x_218); -lean_inc(x_1); -x_220 = l_Lean_Meta_Grind_registerParent(x_1, x_217, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_219); -lean_dec(x_217); -x_221 = lean_ctor_get(x_220, 0); -lean_inc(x_221); -x_222 = lean_ctor_get(x_220, 1); -lean_inc(x_222); -lean_dec(x_220); -x_223 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_221, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_222); -lean_dec(x_221); -return x_223; } else { -uint8_t x_224; -lean_dec(x_217); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +uint8_t x_30; lean_dec(x_2); -lean_dec(x_1); -x_224 = !lean_is_exclusive(x_218); -if (x_224 == 0) +x_30 = !lean_is_exclusive(x_22); +if (x_30 == 0) { -return x_218; +return x_22; } else { -lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_225 = lean_ctor_get(x_218, 0); -x_226 = lean_ctor_get(x_218, 1); -lean_inc(x_226); -lean_inc(x_225); -lean_dec(x_218); -x_227 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_227, 0, x_225); -lean_ctor_set(x_227, 1, x_226); -return x_227; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_22, 0); +x_32 = lean_ctor_get(x_22, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_22); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } -else -{ -lean_object* x_228; lean_object* x_229; -x_228 = lean_array_fget(x_4, x_214); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_228); -x_229 = l_Lean_Meta_Grind_internalize(x_228, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_229) == 0) -{ -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; -x_230 = lean_ctor_get(x_229, 1); -lean_inc(x_230); -lean_dec(x_229); -lean_inc(x_1); -x_231 = l_Lean_Meta_Grind_registerParent(x_1, x_228, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_230); -lean_dec(x_228); -x_232 = lean_ctor_get(x_231, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_231, 1); -lean_inc(x_233); -lean_dec(x_231); -x_234 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_232, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_233); -lean_dec(x_232); -return x_234; } -else -{ -uint8_t x_235; -lean_dec(x_228); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_235 = !lean_is_exclusive(x_229); -if (x_235 == 0) -{ -return x_229; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1() { +_start: { -lean_object* x_236; lean_object* x_237; lean_object* x_238; -x_236 = lean_ctor_get(x_229, 0); -x_237 = lean_ctor_get(x_229, 1); -lean_inc(x_237); -lean_inc(x_236); -lean_dec(x_229); -x_238 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_238, 0, x_236); -lean_ctor_set(x_238, 1, x_237); -return x_238; -} -} -} +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } -block_206: -{ -lean_object* x_198; -lean_dec(x_197); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -x_198 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(x_195, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_195); -if (lean_obj_tag(x_198) == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2() { +_start: { -lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_199 = lean_ctor_get(x_198, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_198, 1); -lean_inc(x_200); -lean_dec(x_198); -x_201 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_196, x_199, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_200); -lean_dec(x_199); -lean_dec(x_4); -return x_201; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__2), 6, 0); +return x_1; } -else -{ -uint8_t x_202; -lean_dec(x_196); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_202 = !lean_is_exclusive(x_198); -if (x_202 == 0) -{ -return x_198; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3() { +_start: { -lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_203 = lean_ctor_get(x_198, 0); -x_204 = lean_ctor_get(x_198, 1); -lean_inc(x_204); -lean_inc(x_203); -lean_dec(x_198); -x_205 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_204); -return x_205; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed), 6, 0); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed), 4, 0); +return x_1; } } -case 5: +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5() { +_start: { -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_239 = lean_ctor_get(x_3, 0); -lean_inc(x_239); -x_240 = lean_ctor_get(x_3, 1); -lean_inc(x_240); -lean_dec(x_3); -x_241 = lean_array_set(x_4, x_5, x_240); -x_242 = lean_unsigned_to_nat(1u); -x_243 = lean_nat_sub(x_5, x_242); -lean_dec(x_5); -x_3 = x_239; -x_4 = x_241; -x_5 = x_243; -goto _start; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed), 4, 0); +return x_1; } -case 6: +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_245; lean_object* x_246; lean_object* x_258; uint8_t x_259; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_245 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_245, 0, x_1); -lean_closure_set(x_245, 1, x_2); -x_258 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_259 = l_Lean_Expr_isConstOf(x_3, x_258); -if (x_259 == 0) +uint8_t x_12; +x_12 = l_Lean_Expr_isBVar(x_1); +if (x_12 == 0) { -lean_object* x_260; -x_260 = lean_box(0); -x_246 = x_260; -goto block_257; -} -else +lean_object* x_13; uint8_t x_14; +x_13 = l___private_Lean_Meta_Tactic_Grind_EMatchTheorem_0__Lean_Meta_Grind_dontCare; +x_14 = lean_expr_eqv(x_1, x_13); +if (x_14 == 0) { -lean_object* x_261; lean_object* x_262; uint8_t x_263; -x_261 = lean_array_get_size(x_4); -x_262 = lean_unsigned_to_nat(2u); -x_263 = lean_nat_dec_eq(x_261, x_262); -if (x_263 == 0) +lean_object* x_15; +x_15 = l_Lean_Meta_Grind_groundPattern_x3f(x_1); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_264; -lean_dec(x_261); -x_264 = lean_box(0); -x_246 = x_264; -goto block_257; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_16); +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; +lean_inc(x_17); +x_19 = lean_mk_array(x_17, x_18); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_17, x_20); +lean_dec(x_17); +x_22 = l_Lean_Expr_withAppAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___spec__2(x_2, x_1, x_19, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_22; } else { -lean_object* x_265; uint8_t x_266; -lean_dec(x_245); -lean_dec(x_3); -x_265 = lean_unsigned_to_nat(0u); -x_266 = lean_nat_dec_lt(x_265, x_261); -lean_dec(x_261); -if (x_266 == 0) -{ -lean_object* x_267; lean_object* x_268; lean_object* x_269; -lean_dec(x_4); -x_267 = l_Lean_instInhabitedExpr; -x_268 = l_outOfBounds___rarg(x_267); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_1); +x_23 = lean_ctor_get(x_15, 0); +lean_inc(x_23); +lean_dec(x_15); +x_24 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2; +x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__3; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_268); -x_269 = l_Lean_Meta_Grind_internalize(x_268, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_269) == 0) -{ -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; -x_270 = lean_ctor_get(x_269, 1); -lean_inc(x_270); -lean_dec(x_269); -lean_inc(x_1); -x_271 = l_Lean_Meta_Grind_registerParent(x_1, x_268, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_270); -lean_dec(x_268); -x_272 = lean_ctor_get(x_271, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_271, 1); -lean_inc(x_273); -lean_dec(x_271); -x_274 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_272, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_273); -lean_dec(x_272); -return x_274; -} -else -{ -uint8_t x_275; -lean_dec(x_268); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_275 = !lean_is_exclusive(x_269); -if (x_275 == 0) -{ -return x_269; -} -else +x_26 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_23, x_24, x_25, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_276; lean_object* x_277; lean_object* x_278; -x_276 = lean_ctor_get(x_269, 0); -x_277 = lean_ctor_get(x_269, 1); -lean_inc(x_277); -lean_inc(x_276); -lean_dec(x_269); -x_278 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_278, 0, x_276); -lean_ctor_set(x_278, 1, x_277); -return x_278; -} -} -} -else +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__4; +x_30 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__5; +lean_inc(x_10); +lean_inc(x_9); +x_31 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_27, x_29, x_30, x_9, x_10, x_28); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_279; lean_object* x_280; -x_279 = lean_array_fget(x_4, x_265); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_279); -x_280 = l_Lean_Meta_Grind_internalize(x_279, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_280) == 0) +x_34 = l_Lean_Meta_Grind_canon(x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; -x_281 = lean_ctor_get(x_280, 1); -lean_inc(x_281); -lean_dec(x_280); -lean_inc(x_1); -x_282 = l_Lean_Meta_Grind_registerParent(x_1, x_279, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_281); -lean_dec(x_279); -x_283 = lean_ctor_get(x_282, 0); -lean_inc(x_283); -x_284 = lean_ctor_get(x_282, 1); -lean_inc(x_284); -lean_dec(x_282); -x_285 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_283, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_284); -lean_dec(x_283); -return x_285; -} -else +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Meta_Grind_shareCommon(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +lean_inc(x_38); +x_40 = l_Lean_Meta_Grind_internalize(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_39); +if (lean_obj_tag(x_40) == 0) { -uint8_t x_286; -lean_dec(x_279); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_286 = !lean_is_exclusive(x_280); -if (x_286 == 0) +uint8_t x_41; +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -return x_280; +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +x_43 = l_Lean_Meta_Grind_mkGroundPattern(x_38); +lean_ctor_set(x_40, 0, x_43); +return x_40; } else { -lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_287 = lean_ctor_get(x_280, 0); -x_288 = lean_ctor_get(x_280, 1); -lean_inc(x_288); -lean_inc(x_287); -lean_dec(x_280); -x_289 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_289, 0, x_287); -lean_ctor_set(x_289, 1, x_288); -return x_289; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_40, 1); +lean_inc(x_44); +lean_dec(x_40); +x_45 = l_Lean_Meta_Grind_mkGroundPattern(x_38); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +return x_46; } } +else +{ +uint8_t x_47; +lean_dec(x_38); +x_47 = !lean_is_exclusive(x_40); +if (x_47 == 0) +{ +return x_40; } +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_40, 0); +x_49 = lean_ctor_get(x_40, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_40); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; } } -block_257: -{ -lean_object* x_247; -lean_dec(x_246); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_247 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_247) == 0) -{ -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; -x_248 = lean_ctor_get(x_247, 1); -lean_inc(x_248); -lean_dec(x_247); -lean_inc(x_1); -x_249 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_248); -lean_dec(x_3); -x_250 = lean_ctor_get(x_249, 0); -lean_inc(x_250); -x_251 = lean_ctor_get(x_249, 1); -lean_inc(x_251); -lean_dec(x_249); -x_252 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_245, x_250, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_251); -lean_dec(x_250); -lean_dec(x_4); -return x_252; } else { -uint8_t x_253; -lean_dec(x_245); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_51; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_253 = !lean_is_exclusive(x_247); -if (x_253 == 0) +x_51 = !lean_is_exclusive(x_34); +if (x_51 == 0) { -return x_247; +return x_34; } else { -lean_object* x_254; lean_object* x_255; lean_object* x_256; -x_254 = lean_ctor_get(x_247, 0); -x_255 = lean_ctor_get(x_247, 1); -lean_inc(x_255); -lean_inc(x_254); -lean_dec(x_247); -x_256 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_256, 0, x_254); -lean_ctor_set(x_256, 1, x_255); -return x_256; -} +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_34, 0); +x_53 = lean_ctor_get(x_34, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_34); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } } -case 7: +else { -lean_object* x_290; lean_object* x_291; lean_object* x_303; uint8_t x_304; +uint8_t x_55; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_290 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_290, 0, x_1); -lean_closure_set(x_290, 1, x_2); -x_303 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_304 = l_Lean_Expr_isConstOf(x_3, x_303); -if (x_304 == 0) +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_55 = !lean_is_exclusive(x_31); +if (x_55 == 0) { -lean_object* x_305; -x_305 = lean_box(0); -x_291 = x_305; -goto block_302; +return x_31; } else { -lean_object* x_306; lean_object* x_307; uint8_t x_308; -x_306 = lean_array_get_size(x_4); -x_307 = lean_unsigned_to_nat(2u); -x_308 = lean_nat_dec_eq(x_306, x_307); -if (x_308 == 0) -{ -lean_object* x_309; -lean_dec(x_306); -x_309 = lean_box(0); -x_291 = x_309; -goto block_302; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_31, 0); +x_57 = lean_ctor_get(x_31, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_31); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} } -else -{ -lean_object* x_310; uint8_t x_311; -lean_dec(x_290); -lean_dec(x_3); -x_310 = lean_unsigned_to_nat(0u); -x_311 = lean_nat_dec_lt(x_310, x_306); -lean_dec(x_306); -if (x_311 == 0) -{ -lean_object* x_312; lean_object* x_313; lean_object* x_314; -lean_dec(x_4); -x_312 = l_Lean_instInhabitedExpr; -x_313 = l_outOfBounds___rarg(x_312); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_313); -x_314 = l_Lean_Meta_Grind_internalize(x_313, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_314) == 0) -{ -lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; -x_315 = lean_ctor_get(x_314, 1); -lean_inc(x_315); -lean_dec(x_314); -lean_inc(x_1); -x_316 = l_Lean_Meta_Grind_registerParent(x_1, x_313, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_315); -lean_dec(x_313); -x_317 = lean_ctor_get(x_316, 0); -lean_inc(x_317); -x_318 = lean_ctor_get(x_316, 1); -lean_inc(x_318); -lean_dec(x_316); -x_319 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_317, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_318); -lean_dec(x_317); -return x_319; } else { -uint8_t x_320; -lean_dec(x_313); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +uint8_t x_59; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_320 = !lean_is_exclusive(x_314); -if (x_320 == 0) -{ -return x_314; -} -else +x_59 = !lean_is_exclusive(x_26); +if (x_59 == 0) { -lean_object* x_321; lean_object* x_322; lean_object* x_323; -x_321 = lean_ctor_get(x_314, 0); -x_322 = lean_ctor_get(x_314, 1); -lean_inc(x_322); -lean_inc(x_321); -lean_dec(x_314); -x_323 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_323, 0, x_321); -lean_ctor_set(x_323, 1, x_322); -return x_323; -} -} +return x_26; } else { -lean_object* x_324; lean_object* x_325; -x_324 = lean_array_fget(x_4, x_310); -lean_dec(x_4); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_324); -x_325 = l_Lean_Meta_Grind_internalize(x_324, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_325) == 0) -{ -lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; -x_326 = lean_ctor_get(x_325, 1); -lean_inc(x_326); -lean_dec(x_325); -lean_inc(x_1); -x_327 = l_Lean_Meta_Grind_registerParent(x_1, x_324, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_326); -lean_dec(x_324); -x_328 = lean_ctor_get(x_327, 0); -lean_inc(x_328); -x_329 = lean_ctor_get(x_327, 1); -lean_inc(x_329); -lean_dec(x_327); -x_330 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_328, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_329); -lean_dec(x_328); -return x_330; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_26, 0); +x_61 = lean_ctor_get(x_26, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_26); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} } else { -uint8_t x_331; -lean_dec(x_324); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_63; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_331 = !lean_is_exclusive(x_325); -if (x_331 == 0) -{ -return x_325; -} -else -{ -lean_object* x_332; lean_object* x_333; lean_object* x_334; -x_332 = lean_ctor_get(x_325, 0); -x_333 = lean_ctor_get(x_325, 1); -lean_inc(x_333); -lean_inc(x_332); -lean_dec(x_325); -x_334 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_334, 0, x_332); -lean_ctor_set(x_334, 1, x_333); -return x_334; -} -} -} -} +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_1); +lean_ctor_set(x_63, 1, x_11); +return x_63; } -block_302: -{ -lean_object* x_292; -lean_dec(x_291); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_292 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_292) == 0) -{ -lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; -x_293 = lean_ctor_get(x_292, 1); -lean_inc(x_293); -lean_dec(x_292); -lean_inc(x_1); -x_294 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_293); -lean_dec(x_3); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -x_296 = lean_ctor_get(x_294, 1); -lean_inc(x_296); -lean_dec(x_294); -x_297 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_290, x_295, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_296); -lean_dec(x_295); -lean_dec(x_4); -return x_297; } else { -uint8_t x_298; -lean_dec(x_290); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +lean_object* x_64; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_298 = !lean_is_exclusive(x_292); -if (x_298 == 0) -{ -return x_292; +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_1); +lean_ctor_set(x_64, 1, x_11); +return x_64; } -else -{ -lean_object* x_299; lean_object* x_300; lean_object* x_301; -x_299 = lean_ctor_get(x_292, 0); -x_300 = lean_ctor_get(x_292, 1); -lean_inc(x_300); -lean_inc(x_299); -lean_dec(x_292); -x_301 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_300); -return x_301; } } +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; } } -case 8: -{ -lean_object* x_335; lean_object* x_336; lean_object* x_348; uint8_t x_349; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_335 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_335, 0, x_1); -lean_closure_set(x_335, 1, x_2); -x_348 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_349 = l_Lean_Expr_isConstOf(x_3, x_348); -if (x_349 == 0) +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2() { +_start: { -lean_object* x_350; -x_350 = lean_box(0); -x_336 = x_350; -goto block_347; +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__1; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; } -else -{ -lean_object* x_351; lean_object* x_352; uint8_t x_353; -x_351 = lean_array_get_size(x_4); -x_352 = lean_unsigned_to_nat(2u); -x_353 = lean_nat_dec_eq(x_351, x_352); -if (x_353 == 0) +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3() { +_start: { -lean_object* x_354; -lean_dec(x_351); -x_354 = lean_box(0); -x_336 = x_354; -goto block_347; +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__2; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; } -else +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4() { +_start: { -lean_object* x_355; uint8_t x_356; -lean_dec(x_335); -lean_dec(x_3); -x_355 = lean_unsigned_to_nat(0u); -x_356 = lean_nat_dec_lt(x_355, x_351); -lean_dec(x_351); -if (x_356 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__3; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5() { +_start: { -lean_object* x_357; lean_object* x_358; lean_object* x_359; -lean_dec(x_4); -x_357 = l_Lean_instInhabitedExpr; -x_358 = l_outOfBounds___rarg(x_357); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_358); -x_359 = l_Lean_Meta_Grind_internalize(x_358, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_359) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__4; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_internalize___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; -x_360 = lean_ctor_get(x_359, 1); -lean_inc(x_360); -lean_dec(x_359); -lean_inc(x_1); -x_361 = l_Lean_Meta_Grind_registerParent(x_1, x_358, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_360); -lean_dec(x_358); -x_362 = lean_ctor_get(x_361, 0); -lean_inc(x_362); -x_363 = lean_ctor_get(x_361, 1); -lean_inc(x_363); -lean_dec(x_361); -x_364 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_362, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_363); -lean_dec(x_362); -return x_364; +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at_Lean_Meta_Grind_internalize___spec__1___closed__5; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; } -else +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +_start: { -uint8_t x_365; -lean_dec(x_358); +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_5, 1); +x_20 = lean_nat_dec_lt(x_7, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_365 = !lean_is_exclusive(x_359); -if (x_365 == 0) -{ -return x_359; -} -else -{ -lean_object* x_366; lean_object* x_367; lean_object* x_368; -x_366 = lean_ctor_get(x_359, 0); -x_367 = lean_ctor_get(x_359, 1); -lean_inc(x_367); -lean_inc(x_366); -lean_dec(x_359); -x_368 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_368, 0, x_366); -lean_ctor_set(x_368, 1, x_367); -return x_368; -} -} +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_6); +lean_ctor_set(x_21, 1, x_18); +return x_21; } else { -lean_object* x_369; lean_object* x_370; -x_369 = lean_array_fget(x_4, x_355); -lean_dec(x_4); +lean_object* x_22; lean_object* x_23; +lean_dec(x_6); +x_22 = lean_array_fget(x_3, x_7); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); lean_inc(x_2); -lean_inc(x_369); -x_370 = l_Lean_Meta_Grind_internalize(x_369, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_370) == 0) +lean_inc(x_22); +x_23 = l_Lean_Meta_Grind_internalize(x_22, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -x_371 = lean_ctor_get(x_370, 1); -lean_inc(x_371); -lean_dec(x_370); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); lean_inc(x_1); -x_372 = l_Lean_Meta_Grind_registerParent(x_1, x_369, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_371); -lean_dec(x_369); -x_373 = lean_ctor_get(x_372, 0); -lean_inc(x_373); -x_374 = lean_ctor_get(x_372, 1); -lean_inc(x_374); -lean_dec(x_372); -x_375 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_373, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_374); -lean_dec(x_373); -return x_375; +x_25 = l_Lean_Meta_Grind_registerParent(x_1, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_24); +lean_dec(x_22); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_ctor_get(x_5, 2); +x_28 = lean_nat_add(x_7, x_27); +lean_dec(x_7); +x_29 = lean_box(0); +x_6 = x_29; +x_7 = x_28; +x_8 = lean_box(0); +x_9 = lean_box(0); +x_18 = x_26; +goto _start; } else { -uint8_t x_376; -lean_dec(x_369); +uint8_t x_31; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_376 = !lean_is_exclusive(x_370); -if (x_376 == 0) +x_31 = !lean_is_exclusive(x_23); +if (x_31 == 0) { -return x_370; +return x_23; } else { -lean_object* x_377; lean_object* x_378; lean_object* x_379; -x_377 = lean_ctor_get(x_370, 0); -x_378 = lean_ctor_get(x_370, 1); -lean_inc(x_378); -lean_inc(x_377); -lean_dec(x_370); -x_379 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_379, 0, x_377); -lean_ctor_set(x_379, 1, x_378); -return x_379; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_23, 0); +x_33 = lean_ctor_get(x_23, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_23); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } } } -block_347: +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_337; -lean_dec(x_336); -lean_inc(x_13); -lean_inc(x_12); +lean_object* x_13; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_337 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_337) == 0) -{ -lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; -x_338 = lean_ctor_get(x_337, 1); -lean_inc(x_338); -lean_dec(x_337); lean_inc(x_1); -x_339 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_338); -lean_dec(x_3); -x_340 = lean_ctor_get(x_339, 0); -lean_inc(x_340); -x_341 = lean_ctor_get(x_339, 1); -lean_inc(x_341); -lean_dec(x_339); -x_342 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_335, x_340, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_341); -lean_dec(x_340); -lean_dec(x_4); -return x_342; -} -else +x_13 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_343; -lean_dec(x_335); +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); lean_dec(x_13); -lean_dec(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_15 = l_Lean_Meta_Grind_addCongrTable(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +lean_inc(x_1); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_updateAppMap(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_19; +} +else +{ +uint8_t x_20; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_343 = !lean_is_exclusive(x_337); -if (x_343 == 0) +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) { -return x_337; +return x_15; } else { -lean_object* x_344; lean_object* x_345; lean_object* x_346; -x_344 = lean_ctor_get(x_337, 0); -x_345 = lean_ctor_get(x_337, 1); -lean_inc(x_345); -lean_inc(x_344); -lean_dec(x_337); -x_346 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_346, 0, x_344); -lean_ctor_set(x_346, 1, x_345); -return x_346; -} -} -} -} -case 9: -{ -lean_object* x_380; lean_object* x_381; lean_object* x_393; uint8_t x_394; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_380 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_380, 0, x_1); -lean_closure_set(x_380, 1, x_2); -x_393 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_394 = l_Lean_Expr_isConstOf(x_3, x_393); -if (x_394 == 0) -{ -lean_object* x_395; -x_395 = lean_box(0); -x_381 = x_395; -goto block_392; +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_396; lean_object* x_397; uint8_t x_398; -x_396 = lean_array_get_size(x_4); -x_397 = lean_unsigned_to_nat(2u); -x_398 = lean_nat_dec_eq(x_396, x_397); -if (x_398 == 0) -{ -lean_object* x_399; -lean_dec(x_396); -x_399 = lean_box(0); -x_381 = x_399; -goto block_392; } -else -{ -lean_object* x_400; uint8_t x_401; -lean_dec(x_380); -lean_dec(x_3); -x_400 = lean_unsigned_to_nat(0u); -x_401 = lean_nat_dec_lt(x_400, x_396); -lean_dec(x_396); -if (x_401 == 0) -{ -lean_object* x_402; lean_object* x_403; lean_object* x_404; -lean_dec(x_4); -x_402 = l_Lean_instInhabitedExpr; -x_403 = l_outOfBounds___rarg(x_402); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_403); -x_404 = l_Lean_Meta_Grind_internalize(x_403, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_404) == 0) -{ -lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; -x_405 = lean_ctor_get(x_404, 1); -lean_inc(x_405); -lean_dec(x_404); -lean_inc(x_1); -x_406 = l_Lean_Meta_Grind_registerParent(x_1, x_403, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_405); -lean_dec(x_403); -x_407 = lean_ctor_get(x_406, 0); -lean_inc(x_407); -x_408 = lean_ctor_get(x_406, 1); -lean_inc(x_408); -lean_dec(x_406); -x_409 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_407, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_408); -lean_dec(x_407); -return x_409; } else { -uint8_t x_410; -lean_dec(x_403); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_24; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_1); -x_410 = !lean_is_exclusive(x_404); -if (x_410 == 0) +x_24 = !lean_is_exclusive(x_13); +if (x_24 == 0) { -return x_404; +return x_13; } else { -lean_object* x_411; lean_object* x_412; lean_object* x_413; -x_411 = lean_ctor_get(x_404, 0); -x_412 = lean_ctor_get(x_404, 1); -lean_inc(x_412); -lean_inc(x_411); -lean_dec(x_404); -x_413 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_413, 0, x_411); -lean_ctor_set(x_413, 1, x_412); -return x_413; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_414; lean_object* x_415; -x_414 = lean_array_fget(x_4, x_400); -lean_dec(x_4); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_15 = lean_array_get_size(x_1); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_15); +lean_ctor_set(x_18, 2, x_17); +x_19 = lean_box(0); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -6969,31 +6568,20 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_414); -x_415 = l_Lean_Meta_Grind_internalize(x_414, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_415) == 0) +x_20 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_internalize___spec__2(x_2, x_3, x_1, x_18, x_18, x_19, x_16, lean_box(0), lean_box(0), x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_18); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; -x_416 = lean_ctor_get(x_415, 1); -lean_inc(x_416); -lean_dec(x_415); -lean_inc(x_1); -x_417 = l_Lean_Meta_Grind_registerParent(x_1, x_414, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_416); -lean_dec(x_414); -x_418 = lean_ctor_get(x_417, 0); -lean_inc(x_418); -x_419 = lean_ctor_get(x_417, 1); -lean_inc(x_419); -lean_dec(x_417); -x_420 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_418, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_419); -lean_dec(x_418); -return x_420; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_apply_10(x_4, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_22; } else { -uint8_t x_421; -lean_dec(x_414); +uint8_t x_23; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7002,34 +6590,72 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_421 = !lean_is_exclusive(x_415); -if (x_421 == 0) +lean_dec(x_4); +x_23 = !lean_is_exclusive(x_20); +if (x_23 == 0) { -return x_415; +return x_20; } else { -lean_object* x_422; lean_object* x_423; lean_object* x_424; -x_422 = lean_ctor_get(x_415, 0); -x_423 = lean_ctor_get(x_415, 1); -lean_inc(x_423); -lean_inc(x_422); -lean_dec(x_415); -x_424 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_424, 0, x_422); -lean_ctor_set(x_424, 1, x_423); -return x_424; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_20, 0); +x_25 = lean_ctor_get(x_20, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_20); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} } } } +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); +return x_1; } } -block_392: +static lean_object* _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2() { +_start: { -lean_object* x_382; -lean_dec(x_381); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2; +x_3 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +switch (lean_obj_tag(x_3)) { +case 0: +{ +lean_object* x_15; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_15 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +lean_inc(x_1); +x_17 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -7040,109 +6666,55 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_382 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_382) == 0) -{ -lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; -x_383 = lean_ctor_get(x_382, 1); -lean_inc(x_383); -lean_dec(x_382); -lean_inc(x_1); -x_384 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_383); -lean_dec(x_3); -x_385 = lean_ctor_get(x_384, 0); -lean_inc(x_385); -x_386 = lean_ctor_get(x_384, 1); -lean_inc(x_386); -lean_dec(x_384); -x_387 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_380, x_385, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_386); -lean_dec(x_385); -lean_dec(x_4); -return x_387; -} -else -{ -uint8_t x_388; -lean_dec(x_380); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_388 = !lean_is_exclusive(x_382); -if (x_388 == 0) -{ -return x_382; -} -else -{ -lean_object* x_389; lean_object* x_390; lean_object* x_391; -x_389 = lean_ctor_get(x_382, 0); -x_390 = lean_ctor_get(x_382, 1); -lean_inc(x_390); -lean_inc(x_389); -lean_dec(x_382); -x_391 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_391, 0, x_389); -lean_ctor_set(x_391, 1, x_390); -return x_391; -} -} -} -} -case 10: +x_19 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_425; lean_object* x_426; lean_object* x_438; uint8_t x_439; -lean_dec(x_5); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_34; uint8_t x_35; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); lean_inc(x_2); lean_inc(x_1); -x_425 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_425, 0, x_1); -lean_closure_set(x_425, 1, x_2); -x_438 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_439 = l_Lean_Expr_isConstOf(x_3, x_438); -if (x_439 == 0) +x_21 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_21, 0, x_1); +lean_closure_set(x_21, 1, x_2); +x_34 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_35 = l_Lean_Expr_isConstOf(x_3, x_34); +if (x_35 == 0) { -lean_object* x_440; -x_440 = lean_box(0); -x_426 = x_440; -goto block_437; +lean_object* x_36; +x_36 = lean_box(0); +x_22 = x_36; +goto block_33; } else { -lean_object* x_441; lean_object* x_442; uint8_t x_443; -x_441 = lean_array_get_size(x_4); -x_442 = lean_unsigned_to_nat(2u); -x_443 = lean_nat_dec_eq(x_441, x_442); -if (x_443 == 0) +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = lean_array_get_size(x_4); +x_38 = lean_unsigned_to_nat(2u); +x_39 = lean_nat_dec_eq(x_37, x_38); +if (x_39 == 0) { -lean_object* x_444; -lean_dec(x_441); -x_444 = lean_box(0); -x_426 = x_444; -goto block_437; +lean_object* x_40; +lean_dec(x_37); +x_40 = lean_box(0); +x_22 = x_40; +goto block_33; } else { -lean_object* x_445; uint8_t x_446; -lean_dec(x_425); +lean_object* x_41; uint8_t x_42; +lean_dec(x_21); lean_dec(x_3); -x_445 = lean_unsigned_to_nat(0u); -x_446 = lean_nat_dec_lt(x_445, x_441); -lean_dec(x_441); -if (x_446 == 0) +x_41 = lean_unsigned_to_nat(0u); +x_42 = lean_nat_dec_lt(x_41, x_37); +lean_dec(x_37); +if (x_42 == 0) { -lean_object* x_447; lean_object* x_448; lean_object* x_449; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_dec(x_4); -x_447 = l_Lean_instInhabitedExpr; -x_448 = l_outOfBounds___rarg(x_447); +x_43 = l_Lean_instInhabitedExpr; +x_44 = l_outOfBounds___rarg(x_43); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -7152,30 +6724,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_448); -x_449 = l_Lean_Meta_Grind_internalize(x_448, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_449) == 0) +lean_inc(x_44); +x_45 = l_Lean_Meta_Grind_internalize(x_44, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; -x_450 = lean_ctor_get(x_449, 1); -lean_inc(x_450); -lean_dec(x_449); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); lean_inc(x_1); -x_451 = l_Lean_Meta_Grind_registerParent(x_1, x_448, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_450); -lean_dec(x_448); -x_452 = lean_ctor_get(x_451, 0); -lean_inc(x_452); -x_453 = lean_ctor_get(x_451, 1); -lean_inc(x_453); -lean_dec(x_451); -x_454 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_452, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_453); -lean_dec(x_452); -return x_454; +x_47 = l_Lean_Meta_Grind_registerParent(x_1, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_46); +lean_dec(x_44); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_48, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_49); +lean_dec(x_48); +return x_50; } else { -uint8_t x_455; -lean_dec(x_448); +uint8_t x_51; +lean_dec(x_44); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7186,30 +6758,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_455 = !lean_is_exclusive(x_449); -if (x_455 == 0) +x_51 = !lean_is_exclusive(x_45); +if (x_51 == 0) { -return x_449; +return x_45; } else { -lean_object* x_456; lean_object* x_457; lean_object* x_458; -x_456 = lean_ctor_get(x_449, 0); -x_457 = lean_ctor_get(x_449, 1); -lean_inc(x_457); -lean_inc(x_456); -lean_dec(x_449); -x_458 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_458, 0, x_456); -lean_ctor_set(x_458, 1, x_457); -return x_458; +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_45, 0); +x_53 = lean_ctor_get(x_45, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_45); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } } else { -lean_object* x_459; lean_object* x_460; -x_459 = lean_array_fget(x_4, x_445); +lean_object* x_55; lean_object* x_56; +x_55 = lean_array_fget(x_4, x_41); lean_dec(x_4); lean_inc(x_13); lean_inc(x_12); @@ -7220,30 +6792,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_459); -x_460 = l_Lean_Meta_Grind_internalize(x_459, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_460) == 0) -{ -lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; -x_461 = lean_ctor_get(x_460, 1); -lean_inc(x_461); -lean_dec(x_460); +lean_inc(x_55); +x_56 = l_Lean_Meta_Grind_internalize(x_55, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); lean_inc(x_1); -x_462 = l_Lean_Meta_Grind_registerParent(x_1, x_459, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_461); -lean_dec(x_459); -x_463 = lean_ctor_get(x_462, 0); -lean_inc(x_463); -x_464 = lean_ctor_get(x_462, 1); -lean_inc(x_464); -lean_dec(x_462); -x_465 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_463, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_464); -lean_dec(x_463); -return x_465; +x_58 = l_Lean_Meta_Grind_registerParent(x_1, x_55, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_57); +lean_dec(x_55); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_60); +lean_dec(x_59); +return x_61; } else { -uint8_t x_466; -lean_dec(x_459); +uint8_t x_62; +lean_dec(x_55); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7254,32 +6826,32 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); -x_466 = !lean_is_exclusive(x_460); -if (x_466 == 0) +x_62 = !lean_is_exclusive(x_56); +if (x_62 == 0) { -return x_460; +return x_56; } else { -lean_object* x_467; lean_object* x_468; lean_object* x_469; -x_467 = lean_ctor_get(x_460, 0); -x_468 = lean_ctor_get(x_460, 1); -lean_inc(x_468); -lean_inc(x_467); -lean_dec(x_460); -x_469 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_469, 0, x_467); -lean_ctor_set(x_469, 1, x_468); -return x_469; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_56, 0); +x_64 = lean_ctor_get(x_56, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_56); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } } } -block_437: +block_33: { -lean_object* x_427; -lean_dec(x_426); +lean_object* x_23; +lean_dec(x_22); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -7290,30 +6862,30 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); lean_inc(x_3); -x_427 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_427) == 0) +x_23 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; -x_428 = lean_ctor_get(x_427, 1); -lean_inc(x_428); -lean_dec(x_427); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); lean_inc(x_1); -x_429 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_428); +x_25 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); lean_dec(x_3); -x_430 = lean_ctor_get(x_429, 0); -lean_inc(x_430); -x_431 = lean_ctor_get(x_429, 1); -lean_inc(x_431); -lean_dec(x_429); -x_432 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_425, x_430, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_431); -lean_dec(x_430); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_21, x_26, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_27); +lean_dec(x_26); lean_dec(x_4); -return x_432; +return x_28; } else { -uint8_t x_433; -lean_dec(x_425); +uint8_t x_29; +lean_dec(x_21); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7326,106 +6898,65 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_433 = !lean_is_exclusive(x_427); -if (x_433 == 0) +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) { -return x_427; +return x_23; } else { -lean_object* x_434; lean_object* x_435; lean_object* x_436; -x_434 = lean_ctor_get(x_427, 0); -x_435 = lean_ctor_get(x_427, 1); -lean_inc(x_435); -lean_inc(x_434); -lean_dec(x_427); -x_436 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_436, 0, x_434); -lean_ctor_set(x_436, 1, x_435); -return x_436; -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } -default: -{ -lean_object* x_470; lean_object* x_471; lean_object* x_483; uint8_t x_484; -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_470 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); -lean_closure_set(x_470, 0, x_1); -lean_closure_set(x_470, 1, x_2); -x_483 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4; -x_484 = l_Lean_Expr_isConstOf(x_3, x_483); -if (x_484 == 0) -{ -lean_object* x_485; -x_485 = lean_box(0); -x_471 = x_485; -goto block_482; } else { -lean_object* x_486; lean_object* x_487; uint8_t x_488; -x_486 = lean_array_get_size(x_4); -x_487 = lean_unsigned_to_nat(2u); -x_488 = lean_nat_dec_eq(x_486, x_487); -if (x_488 == 0) +uint8_t x_66; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_19); +if (x_66 == 0) { -lean_object* x_489; -lean_dec(x_486); -x_489 = lean_box(0); -x_471 = x_489; -goto block_482; +return x_19; } else { -lean_object* x_490; uint8_t x_491; -lean_dec(x_470); -lean_dec(x_3); -x_490 = lean_unsigned_to_nat(0u); -x_491 = lean_nat_dec_lt(x_490, x_486); -lean_dec(x_486); -if (x_491 == 0) -{ -lean_object* x_492; lean_object* x_493; lean_object* x_494; -lean_dec(x_4); -x_492 = l_Lean_instInhabitedExpr; -x_493 = l_outOfBounds___rarg(x_492); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_493); -x_494 = l_Lean_Meta_Grind_internalize(x_493, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_494) == 0) -{ -lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; -x_495 = lean_ctor_get(x_494, 1); -lean_inc(x_495); -lean_dec(x_494); -lean_inc(x_1); -x_496 = l_Lean_Meta_Grind_registerParent(x_1, x_493, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_495); -lean_dec(x_493); -x_497 = lean_ctor_get(x_496, 0); -lean_inc(x_497); -x_498 = lean_ctor_get(x_496, 1); -lean_inc(x_498); -lean_dec(x_496); -x_499 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_497, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_498); -lean_dec(x_497); -return x_499; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_19, 0); +x_68 = lean_ctor_get(x_19, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_19); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} } else { -uint8_t x_500; -lean_dec(x_493); +uint8_t x_70; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7434,33 +6965,51 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_500 = !lean_is_exclusive(x_494); -if (x_500 == 0) +x_70 = !lean_is_exclusive(x_15); +if (x_70 == 0) { -return x_494; +return x_15; } else { -lean_object* x_501; lean_object* x_502; lean_object* x_503; -x_501 = lean_ctor_get(x_494, 0); -x_502 = lean_ctor_get(x_494, 1); -lean_inc(x_502); -lean_inc(x_501); -lean_dec(x_494); -x_503 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_503, 0, x_501); -lean_ctor_set(x_503, 1, x_502); -return x_503; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_15, 0); +x_72 = lean_ctor_get(x_15, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_15); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +return x_73; } } } -else +case 1: { -lean_object* x_504; lean_object* x_505; -x_504 = lean_array_fget(x_4, x_490); -lean_dec(x_4); +lean_object* x_74; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_74 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +lean_inc(x_1); +x_76 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_75); +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +lean_dec(x_76); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -7470,66 +7019,56 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_504); -x_505 = l_Lean_Meta_Grind_internalize(x_504, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_505) == 0) +lean_inc(x_3); +x_78 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_77); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; -x_506 = lean_ctor_get(x_505, 1); -lean_inc(x_506); -lean_dec(x_505); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_93; uint8_t x_94; +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +lean_inc(x_2); lean_inc(x_1); -x_507 = l_Lean_Meta_Grind_registerParent(x_1, x_504, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_506); -lean_dec(x_504); -x_508 = lean_ctor_get(x_507, 0); -lean_inc(x_508); -x_509 = lean_ctor_get(x_507, 1); -lean_inc(x_509); -lean_dec(x_507); -x_510 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_508, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_509); -lean_dec(x_508); -return x_510; +x_80 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_80, 0, x_1); +lean_closure_set(x_80, 1, x_2); +x_93 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_94 = l_Lean_Expr_isConstOf(x_3, x_93); +if (x_94 == 0) +{ +lean_object* x_95; +x_95 = lean_box(0); +x_81 = x_95; +goto block_92; } else { -uint8_t x_511; -lean_dec(x_504); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -lean_dec(x_1); -x_511 = !lean_is_exclusive(x_505); -if (x_511 == 0) +lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_96 = lean_array_get_size(x_4); +x_97 = lean_unsigned_to_nat(2u); +x_98 = lean_nat_dec_eq(x_96, x_97); +if (x_98 == 0) { -return x_505; +lean_object* x_99; +lean_dec(x_96); +x_99 = lean_box(0); +x_81 = x_99; +goto block_92; } else { -lean_object* x_512; lean_object* x_513; lean_object* x_514; -x_512 = lean_ctor_get(x_505, 0); -x_513 = lean_ctor_get(x_505, 1); -lean_inc(x_513); -lean_inc(x_512); -lean_dec(x_505); -x_514 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_514, 0, x_512); -lean_ctor_set(x_514, 1, x_513); -return x_514; -} -} -} -} -} -block_482: +lean_object* x_100; uint8_t x_101; +lean_dec(x_80); +lean_dec(x_3); +x_100 = lean_unsigned_to_nat(0u); +x_101 = lean_nat_dec_lt(x_100, x_96); +lean_dec(x_96); +if (x_101 == 0) { -lean_object* x_472; -lean_dec(x_471); +lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_4); +x_102 = l_Lean_instInhabitedExpr; +x_103 = l_outOfBounds___rarg(x_102); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -7539,31 +7078,30 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_2); -lean_inc(x_3); -x_472 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_472) == 0) +lean_inc(x_103); +x_104 = l_Lean_Meta_Grind_internalize(x_103, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_79); +if (lean_obj_tag(x_104) == 0) { -lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; -x_473 = lean_ctor_get(x_472, 1); -lean_inc(x_473); -lean_dec(x_472); +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +lean_dec(x_104); lean_inc(x_1); -x_474 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_473); -lean_dec(x_3); -x_475 = lean_ctor_get(x_474, 0); -lean_inc(x_475); -x_476 = lean_ctor_get(x_474, 1); -lean_inc(x_476); -lean_dec(x_474); -x_477 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_470, x_475, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_476); -lean_dec(x_475); -lean_dec(x_4); -return x_477; +x_106 = l_Lean_Meta_Grind_registerParent(x_1, x_103, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_105); +lean_dec(x_103); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_107, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_108); +lean_dec(x_107); +return x_109; } else { -uint8_t x_478; -lean_dec(x_470); +uint8_t x_110; +lean_dec(x_103); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7572,2554 +7110,7655 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_478 = !lean_is_exclusive(x_472); -if (x_478 == 0) +x_110 = !lean_is_exclusive(x_104); +if (x_110 == 0) { -return x_472; +return x_104; } else { -lean_object* x_479; lean_object* x_480; lean_object* x_481; -x_479 = lean_ctor_get(x_472, 0); -x_480 = lean_ctor_get(x_472, 1); -lean_inc(x_480); -lean_inc(x_479); -lean_dec(x_472); -x_481 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_481, 0, x_479); -lean_ctor_set(x_481, 1, x_480); -return x_481; -} -} -} -} -} -} +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_104, 0); +x_112 = lean_ctor_get(x_104, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_104); +x_113 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +return x_113; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = 0; -x_14 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_13, x_13, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Internalize", 34, 34); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__2() { -_start: +lean_object* x_114; lean_object* x_115; +x_114 = lean_array_fget(x_4, x_100); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_114); +x_115 = l_Lean_Meta_Grind_internalize(x_114, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_79); +if (lean_obj_tag(x_115) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.internalize", 27, 27); -return x_1; -} +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_116 = lean_ctor_get(x_115, 1); +lean_inc(x_116); +lean_dec(x_115); +lean_inc(x_1); +x_117 = l_Lean_Meta_Grind_registerParent(x_1, x_114, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_116); +lean_dec(x_114); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec(x_117); +x_120 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_118, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_119); +lean_dec(x_118); +return x_120; } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__4() { -_start: +uint8_t x_121; +lean_dec(x_114); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_121 = !lean_is_exclusive(x_115); +if (x_121 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_internalize___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_internalize___lambda__2___closed__2; -x_3 = lean_unsigned_to_nat(76u); -x_4 = lean_unsigned_to_nat(16u); -x_5 = l_Lean_Meta_Grind_internalize___lambda__2___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} +return x_115; } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unexpected term during internalization", 38, 38); -return x_1; +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_115, 0); +x_123 = lean_ctor_get(x_115, 1); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_115); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +return x_124; } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_internalize___lambda__2___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_2); -lean_dec(x_1); -x_13 = l_Lean_Meta_Grind_internalize___lambda__2___closed__4; -x_14 = l_panic___at_Lean_Meta_Grind_internalize___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; } -case 2: +block_92: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_82; +lean_dec(x_81); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_82 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_79); +if (lean_obj_tag(x_82) == 0) { -lean_object* x_19; uint8_t x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_dec(x_16); -x_20 = 0; -x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_20, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +lean_dec(x_82); +lean_inc(x_1); +x_84 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_83); +lean_dec(x_3); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_80, x_85, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_86); +lean_dec(x_85); lean_dec(x_4); -return x_21; +return x_87; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_16); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; -x_23 = lean_ctor_get(x_16, 1); -x_24 = lean_ctor_get(x_16, 0); -lean_dec(x_24); -lean_inc(x_1); -x_25 = l_Lean_indentExpr(x_1); -x_26 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_16, 7); -lean_ctor_set(x_16, 1, x_25); -lean_ctor_set(x_16, 0, x_26); -x_27 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_16); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_15, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = 0; -x_32 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_31, x_31, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +uint8_t x_88; +lean_dec(x_80); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -return x_32; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_88 = !lean_is_exclusive(x_82); +if (x_88 == 0) +{ +return x_82; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; -x_33 = lean_ctor_get(x_16, 1); -lean_inc(x_33); -lean_dec(x_16); -lean_inc(x_1); -x_34 = l_Lean_indentExpr(x_1); -x_35 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -x_37 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_15, x_38, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -x_41 = 0; -x_42 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_41, x_41, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_42; +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_82, 0); +x_90 = lean_ctor_get(x_82, 1); +lean_inc(x_90); +lean_inc(x_89); +lean_dec(x_82); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; } } } -case 3: +} +else { -lean_object* x_43; lean_object* x_44; +uint8_t x_125; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_12); -return x_44; -} -case 4: -{ -lean_object* x_45; -x_45 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_45; -} -case 5: -{ -lean_object* x_46; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_46 = l_Lean_Meta_isLitValue(x_1, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; uint8_t x_48; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_unbox(x_47); -lean_dec(x_47); -if (x_48 == 0) +x_125 = !lean_is_exclusive(x_78); +if (x_125 == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); -x_50 = lean_unsigned_to_nat(0u); -x_51 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_50); -x_52 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; -lean_inc(x_51); -x_53 = lean_mk_array(x_51, x_52); -x_54 = lean_unsigned_to_nat(1u); -x_55 = lean_nat_sub(x_51, x_54); -lean_dec(x_51); -lean_inc(x_1); -x_56 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(x_1, x_2, x_1, x_53, x_55, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_49); -return x_56; +return x_78; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_46, 1); -lean_inc(x_57); -lean_dec(x_46); -x_58 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_57); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_58; +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_78, 0); +x_127 = lean_ctor_get(x_78, 1); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_78); +x_128 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +return x_128; +} } } else { -uint8_t x_59; +uint8_t x_129; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_46); -if (x_59 == 0) +x_129 = !lean_is_exclusive(x_74); +if (x_129 == 0) { -return x_46; +return x_74; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_46, 0); -x_61 = lean_ctor_get(x_46, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_46); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_74, 0); +x_131 = lean_ctor_get(x_74, 1); +lean_inc(x_131); +lean_inc(x_130); +lean_dec(x_74); +x_132 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +return x_132; } } } -case 7: +case 2: { -lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_63 = lean_ctor_get(x_1, 1); -lean_inc(x_63); -x_64 = 0; +lean_object* x_133; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_133 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_133) == 0) +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_134 = lean_ctor_get(x_133, 1); +lean_inc(x_134); +lean_dec(x_133); +lean_inc(x_1); +x_135 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_134); +x_136 = lean_ctor_get(x_135, 1); +lean_inc(x_136); +lean_dec(x_135); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_137 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_136); +if (lean_obj_tag(x_137) == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_152; uint8_t x_153; +x_138 = lean_ctor_get(x_137, 1); +lean_inc(x_138); +lean_dec(x_137); lean_inc(x_2); lean_inc(x_1); -x_65 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_64, x_64, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); +x_139 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_139, 0, x_1); +lean_closure_set(x_139, 1, x_2); +x_152 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_153 = l_Lean_Expr_isConstOf(x_3, x_152); +if (x_153 == 0) +{ +lean_object* x_154; +x_154 = lean_box(0); +x_140 = x_154; +goto block_151; +} +else +{ +lean_object* x_155; lean_object* x_156; uint8_t x_157; +x_155 = lean_array_get_size(x_4); +x_156 = lean_unsigned_to_nat(2u); +x_157 = lean_nat_dec_eq(x_155, x_156); +if (x_157 == 0) +{ +lean_object* x_158; +lean_dec(x_155); +x_158 = lean_box(0); +x_140 = x_158; +goto block_151; +} +else +{ +lean_object* x_159; uint8_t x_160; +lean_dec(x_139); +lean_dec(x_3); +x_159 = lean_unsigned_to_nat(0u); +x_160 = lean_nat_dec_lt(x_159, x_155); +lean_dec(x_155); +if (x_160 == 0) +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_4); +x_161 = l_Lean_instInhabitedExpr; +x_162 = l_outOfBounds___rarg(x_161); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_63); -x_67 = l_Lean_Meta_isProp(x_63, x_8, x_9, x_10, x_11, x_66); -if (lean_obj_tag(x_67) == 0) +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_162); +x_163 = l_Lean_Meta_Grind_internalize(x_162, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +if (lean_obj_tag(x_163) == 0) { -lean_object* x_68; uint8_t x_69; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_unbox(x_68); -lean_dec(x_68); -if (x_69 == 0) +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_164 = lean_ctor_get(x_163, 1); +lean_inc(x_164); +lean_dec(x_163); +lean_inc(x_1); +x_165 = l_Lean_Meta_Grind_registerParent(x_1, x_162, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_164); +lean_dec(x_162); +x_166 = lean_ctor_get(x_165, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_165, 1); +lean_inc(x_167); +lean_dec(x_165); +x_168 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_166, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_167); +lean_dec(x_166); +return x_168; +} +else { -uint8_t x_70; -lean_dec(x_63); +uint8_t x_169; +lean_dec(x_162); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_70 = !lean_is_exclusive(x_67); -if (x_70 == 0) +x_169 = !lean_is_exclusive(x_163); +if (x_169 == 0) { -lean_object* x_71; lean_object* x_72; -x_71 = lean_ctor_get(x_67, 0); -lean_dec(x_71); -x_72 = lean_box(0); -lean_ctor_set(x_67, 0, x_72); -return x_67; +return x_163; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_67, 1); -lean_inc(x_73); -lean_dec(x_67); -x_74 = lean_box(0); -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_73); -return x_75; +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_163, 0); +x_171 = lean_ctor_get(x_163, 1); +lean_inc(x_171); +lean_inc(x_170); +lean_dec(x_163); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +return x_172; +} } } else { -lean_object* x_76; lean_object* x_77; -x_76 = lean_ctor_get(x_67, 1); -lean_inc(x_76); -lean_dec(x_67); +lean_object* x_173; lean_object* x_174; +x_173 = lean_array_fget(x_4, x_159); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_1); -x_77 = l_Lean_Meta_isProp(x_1, x_8, x_9, x_10, x_11, x_76); -if (lean_obj_tag(x_77) == 0) +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_173); +x_174 = l_Lean_Meta_Grind_internalize(x_173, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +if (lean_obj_tag(x_174) == 0) { -lean_object* x_78; uint8_t x_79; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_unbox(x_78); -lean_dec(x_78); -if (x_79 == 0) +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +lean_dec(x_174); +lean_inc(x_1); +x_176 = l_Lean_Meta_Grind_registerParent(x_1, x_173, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_175); +lean_dec(x_173); +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_176, 1); +lean_inc(x_178); +lean_dec(x_176); +x_179 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_177, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_178); +lean_dec(x_177); +return x_179; +} +else { -uint8_t x_80; -lean_dec(x_63); +uint8_t x_180; +lean_dec(x_173); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_80 = !lean_is_exclusive(x_77); -if (x_80 == 0) +x_180 = !lean_is_exclusive(x_174); +if (x_180 == 0) { -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_77, 0); -lean_dec(x_81); -x_82 = lean_box(0); -lean_ctor_set(x_77, 0, x_82); -return x_77; +return x_174; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_77, 1); -lean_inc(x_83); -lean_dec(x_77); -x_84 = lean_box(0); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_83); -return x_85; +lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_181 = lean_ctor_get(x_174, 0); +x_182 = lean_ctor_get(x_174, 1); +lean_inc(x_182); +lean_inc(x_181); +lean_dec(x_174); +x_183 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_183, 0, x_181); +lean_ctor_set(x_183, 1, x_182); +return x_183; } } -else +} +} +} +block_151: { -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_77, 1); -lean_inc(x_86); -lean_dec(x_77); +lean_object* x_141; +lean_dec(x_140); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_63); -x_87 = l_Lean_Meta_Grind_internalize(x_63, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_86); -if (lean_obj_tag(x_87) == 0) +lean_inc(x_2); +lean_inc(x_3); +x_141 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_138); +if (lean_obj_tag(x_141) == 0) { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_142 = lean_ctor_get(x_141, 1); +lean_inc(x_142); +lean_dec(x_141); lean_inc(x_1); -x_89 = l_Lean_Meta_Grind_registerParent(x_1, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_88); -lean_dec(x_63); -x_90 = lean_ctor_get(x_89, 1); -lean_inc(x_90); -lean_dec(x_89); -x_91 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_90); -return x_91; +x_143 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_142); +lean_dec(x_3); +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +x_146 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_139, x_144, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_145); +lean_dec(x_144); +lean_dec(x_4); +return x_146; } else { -uint8_t x_92; -lean_dec(x_63); +uint8_t x_147; +lean_dec(x_139); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_92 = !lean_is_exclusive(x_87); -if (x_92 == 0) +x_147 = !lean_is_exclusive(x_141); +if (x_147 == 0) { -return x_87; +return x_141; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_87, 0); -x_94 = lean_ctor_get(x_87, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_87); -x_95 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_94); -return x_95; +lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_141, 0); +x_149 = lean_ctor_get(x_141, 1); +lean_inc(x_149); +lean_inc(x_148); +lean_dec(x_141); +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +return x_150; } } } } else { -uint8_t x_96; -lean_dec(x_63); +uint8_t x_184; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_96 = !lean_is_exclusive(x_77); -if (x_96 == 0) +x_184 = !lean_is_exclusive(x_137); +if (x_184 == 0) { -return x_77; +return x_137; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_77, 0); -x_98 = lean_ctor_get(x_77, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_77); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -return x_99; -} +lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_185 = lean_ctor_get(x_137, 0); +x_186 = lean_ctor_get(x_137, 1); +lean_inc(x_186); +lean_inc(x_185); +lean_dec(x_137); +x_187 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_187, 0, x_185); +lean_ctor_set(x_187, 1, x_186); +return x_187; } } } else { -uint8_t x_100; -lean_dec(x_63); +uint8_t x_188; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_100 = !lean_is_exclusive(x_67); -if (x_100 == 0) +x_188 = !lean_is_exclusive(x_133); +if (x_188 == 0) { -return x_67; +return x_133; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_67, 0); -x_102 = lean_ctor_get(x_67, 1); -lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_67); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -return x_103; +lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_189 = lean_ctor_get(x_133, 0); +x_190 = lean_ctor_get(x_133, 1); +lean_inc(x_190); +lean_inc(x_189); +lean_dec(x_133); +x_191 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_190); +return x_191; } } } -case 9: +case 3: { -lean_object* x_104; -x_104 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_192; lean_dec(x_5); -lean_dec(x_4); -return x_104; -} -case 10: +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_192 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_192) == 0) { -lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_105 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_106 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_105, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_unbox(x_107); -lean_dec(x_107); -if (x_108 == 0) +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; +x_193 = lean_ctor_get(x_192, 1); +lean_inc(x_193); +lean_dec(x_192); +lean_inc(x_1); +x_194 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_193); +x_195 = lean_ctor_get(x_194, 1); +lean_inc(x_195); +lean_dec(x_194); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_196 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_195); +if (lean_obj_tag(x_196) == 0) { -lean_object* x_109; uint8_t x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_106, 1); -lean_inc(x_109); -lean_dec(x_106); -x_110 = 0; -x_111 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_110, x_110, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_109); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_111; +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_211; uint8_t x_212; +x_197 = lean_ctor_get(x_196, 1); +lean_inc(x_197); +lean_dec(x_196); +lean_inc(x_2); +lean_inc(x_1); +x_198 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_198, 0, x_1); +lean_closure_set(x_198, 1, x_2); +x_211 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_212 = l_Lean_Expr_isConstOf(x_3, x_211); +if (x_212 == 0) +{ +lean_object* x_213; +x_213 = lean_box(0); +x_199 = x_213; +goto block_210; } else { -uint8_t x_112; -x_112 = !lean_is_exclusive(x_106); -if (x_112 == 0) +lean_object* x_214; lean_object* x_215; uint8_t x_216; +x_214 = lean_array_get_size(x_4); +x_215 = lean_unsigned_to_nat(2u); +x_216 = lean_nat_dec_eq(x_214, x_215); +if (x_216 == 0) { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; lean_object* x_122; -x_113 = lean_ctor_get(x_106, 1); -x_114 = lean_ctor_get(x_106, 0); -lean_dec(x_114); -lean_inc(x_1); -x_115 = l_Lean_indentExpr(x_1); -x_116 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_106, 7); -lean_ctor_set(x_106, 1, x_115); -lean_ctor_set(x_106, 0, x_116); -x_117 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_118 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_118, 0, x_106); -lean_ctor_set(x_118, 1, x_117); -x_119 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_105, x_118, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_113); -x_120 = lean_ctor_get(x_119, 1); -lean_inc(x_120); -lean_dec(x_119); -x_121 = 0; -x_122 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_121, x_121, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_120); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_122; +lean_object* x_217; +lean_dec(x_214); +x_217 = lean_box(0); +x_199 = x_217; +goto block_210; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; lean_object* x_132; -x_123 = lean_ctor_get(x_106, 1); -lean_inc(x_123); -lean_dec(x_106); +lean_object* x_218; uint8_t x_219; +lean_dec(x_198); +lean_dec(x_3); +x_218 = lean_unsigned_to_nat(0u); +x_219 = lean_nat_dec_lt(x_218, x_214); +lean_dec(x_214); +if (x_219 == 0) +{ +lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_dec(x_4); +x_220 = l_Lean_instInhabitedExpr; +x_221 = l_outOfBounds___rarg(x_220); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_221); +x_222 = l_Lean_Meta_Grind_internalize(x_221, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_197); +if (lean_obj_tag(x_222) == 0) +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; +x_223 = lean_ctor_get(x_222, 1); +lean_inc(x_223); +lean_dec(x_222); lean_inc(x_1); -x_124 = l_Lean_indentExpr(x_1); -x_125 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_126 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_124); -x_127 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_128 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -x_129 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_105, x_128, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_123); -x_130 = lean_ctor_get(x_129, 1); -lean_inc(x_130); -lean_dec(x_129); -x_131 = 0; -x_132 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_131, x_131, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_130); +x_224 = l_Lean_Meta_Grind_registerParent(x_1, x_221, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_223); +lean_dec(x_221); +x_225 = lean_ctor_get(x_224, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_224, 1); +lean_inc(x_226); +lean_dec(x_224); +x_227 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_225, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_226); +lean_dec(x_225); +return x_227; +} +else +{ +uint8_t x_228; +lean_dec(x_221); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_132; +lean_dec(x_2); +lean_dec(x_1); +x_228 = !lean_is_exclusive(x_222); +if (x_228 == 0) +{ +return x_222; +} +else +{ +lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_229 = lean_ctor_get(x_222, 0); +x_230 = lean_ctor_get(x_222, 1); +lean_inc(x_230); +lean_inc(x_229); +lean_dec(x_222); +x_231 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_231, 0, x_229); +lean_ctor_set(x_231, 1, x_230); +return x_231; } } } -case 11: +else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_133 = l_Lean_Meta_Grind_addCongrTable___closed__2; -x_134 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_133, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_unbox(x_135); -lean_dec(x_135); -if (x_136 == 0) +lean_object* x_232; lean_object* x_233; +x_232 = lean_array_fget(x_4, x_218); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_232); +x_233 = l_Lean_Meta_Grind_internalize(x_232, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_197); +if (lean_obj_tag(x_233) == 0) { -lean_object* x_137; uint8_t x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_134, 1); -lean_inc(x_137); -lean_dec(x_134); -x_138 = 0; -x_139 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_138, x_138, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_137); +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_234 = lean_ctor_get(x_233, 1); +lean_inc(x_234); +lean_dec(x_233); +lean_inc(x_1); +x_235 = l_Lean_Meta_Grind_registerParent(x_1, x_232, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_234); +lean_dec(x_232); +x_236 = lean_ctor_get(x_235, 0); +lean_inc(x_236); +x_237 = lean_ctor_get(x_235, 1); +lean_inc(x_237); +lean_dec(x_235); +x_238 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_236, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_237); +lean_dec(x_236); +return x_238; +} +else +{ +uint8_t x_239; +lean_dec(x_232); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_139; +lean_dec(x_2); +lean_dec(x_1); +x_239 = !lean_is_exclusive(x_233); +if (x_239 == 0) +{ +return x_233; } else { -uint8_t x_140; -x_140 = !lean_is_exclusive(x_134); -if (x_140 == 0) +lean_object* x_240; lean_object* x_241; lean_object* x_242; +x_240 = lean_ctor_get(x_233, 0); +x_241 = lean_ctor_get(x_233, 1); +lean_inc(x_241); +lean_inc(x_240); +lean_dec(x_233); +x_242 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_242, 0, x_240); +lean_ctor_set(x_242, 1, x_241); +return x_242; +} +} +} +} +} +block_210: { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; -x_141 = lean_ctor_get(x_134, 1); -x_142 = lean_ctor_get(x_134, 0); -lean_dec(x_142); +lean_object* x_200; +lean_dec(x_199); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_200 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_197); +if (lean_obj_tag(x_200) == 0) +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_201 = lean_ctor_get(x_200, 1); +lean_inc(x_201); +lean_dec(x_200); lean_inc(x_1); -x_143 = l_Lean_indentExpr(x_1); -x_144 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -lean_ctor_set_tag(x_134, 7); -lean_ctor_set(x_134, 1, x_143); -lean_ctor_set(x_134, 0, x_144); -x_145 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_146 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_146, 0, x_134); -lean_ctor_set(x_146, 1, x_145); -x_147 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_133, x_146, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_141); -x_148 = lean_ctor_get(x_147, 1); -lean_inc(x_148); -lean_dec(x_147); -x_149 = 0; -x_150 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_149, x_149, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_148); +x_202 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_201); +lean_dec(x_3); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_202, 1); +lean_inc(x_204); +lean_dec(x_202); +x_205 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_198, x_203, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_204); +lean_dec(x_203); +lean_dec(x_4); +return x_205; +} +else +{ +uint8_t x_206; +lean_dec(x_198); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -return x_150; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_206 = !lean_is_exclusive(x_200); +if (x_206 == 0) +{ +return x_200; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; lean_object* x_160; -x_151 = lean_ctor_get(x_134, 1); -lean_inc(x_151); -lean_dec(x_134); -lean_inc(x_1); -x_152 = l_Lean_indentExpr(x_1); -x_153 = l_Lean_Meta_Grind_internalize___lambda__2___closed__6; -x_154 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_152); -x_155 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_156 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_156, 0, x_154); -lean_ctor_set(x_156, 1, x_155); -x_157 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_133, x_156, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_151); -x_158 = lean_ctor_get(x_157, 1); -lean_inc(x_158); -lean_dec(x_157); -x_159 = 0; -x_160 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_159, x_159, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_158); +lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_207 = lean_ctor_get(x_200, 0); +x_208 = lean_ctor_get(x_200, 1); +lean_inc(x_208); +lean_inc(x_207); +lean_dec(x_200); +x_209 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set(x_209, 1, x_208); +return x_209; +} +} +} +} +else +{ +uint8_t x_243; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -return x_160; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_243 = !lean_is_exclusive(x_196); +if (x_243 == 0) +{ +return x_196; +} +else +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_244 = lean_ctor_get(x_196, 0); +x_245 = lean_ctor_get(x_196, 1); +lean_inc(x_245); +lean_inc(x_244); +lean_dec(x_196); +x_246 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_246, 0, x_244); +lean_ctor_set(x_246, 1, x_245); +return x_246; } } } -default: +else { -uint8_t x_161; lean_object* x_162; -x_161 = 0; -x_162 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_161, x_161, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +uint8_t x_247; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_4); -return x_162; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_247 = !lean_is_exclusive(x_192); +if (x_247 == 0) +{ +return x_192; } +else +{ +lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_248 = lean_ctor_get(x_192, 0); +x_249 = lean_ctor_get(x_192, 1); +lean_inc(x_249); +lean_inc(x_248); +lean_dec(x_192); +x_250 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_250, 0, x_248); +lean_ctor_set(x_250, 1, x_249); +return x_250; } } } -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__1() { -_start: +case 4: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("internalize", 11, 11); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__2() { -_start: +lean_object* x_251; lean_object* x_252; +lean_dec(x_5); +x_251 = lean_ctor_get(x_3, 0); +lean_inc(x_251); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_252 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_252) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_internalize___lambda__3___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_253 = lean_ctor_get(x_252, 1); +lean_inc(x_253); +lean_dec(x_252); +lean_inc(x_1); +x_254 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_253); +x_255 = lean_ctor_get(x_254, 1); +lean_inc(x_255); +lean_dec(x_254); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_256 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_255); +if (lean_obj_tag(x_256) == 0) +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_269; uint8_t x_270; +x_257 = lean_ctor_get(x_256, 1); +lean_inc(x_257); +lean_dec(x_256); +lean_inc(x_2); +lean_inc(x_1); +x_258 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_258, 0, x_1); +lean_closure_set(x_258, 1, x_2); +x_269 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_270 = l_Lean_Expr_isConstOf(x_3, x_269); +lean_dec(x_3); +if (x_270 == 0) +{ +lean_object* x_271; +x_271 = lean_box(0); +x_259 = x_271; +goto block_268; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = l_Lean_Meta_Grind_internalize___lambda__3___closed__2; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_unbox(x_15); -lean_dec(x_15); -if (x_16 == 0) +lean_object* x_272; lean_object* x_273; uint8_t x_274; +x_272 = lean_array_get_size(x_4); +x_273 = lean_unsigned_to_nat(2u); +x_274 = lean_nat_dec_eq(x_272, x_273); +if (x_274 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_box(0); -x_19 = l_Lean_Meta_Grind_internalize___lambda__2(x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); -return x_19; +lean_object* x_275; +lean_dec(x_272); +x_275 = lean_box(0); +x_259 = x_275; +goto block_268; } else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_14); -if (x_20 == 0) +lean_object* x_276; uint8_t x_277; +lean_dec(x_258); +lean_dec(x_251); +x_276 = lean_unsigned_to_nat(0u); +x_277 = lean_nat_dec_lt(x_276, x_272); +lean_dec(x_272); +if (x_277 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_21 = lean_ctor_get(x_14, 1); -x_22 = lean_ctor_get(x_14, 0); -lean_dec(x_22); +lean_object* x_278; lean_object* x_279; lean_object* x_280; +lean_dec(x_4); +x_278 = l_Lean_instInhabitedExpr; +x_279 = l_outOfBounds___rarg(x_278); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_279); +x_280 = l_Lean_Meta_Grind_internalize(x_279, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_257); +if (lean_obj_tag(x_280) == 0) +{ +lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_281 = lean_ctor_get(x_280, 1); +lean_inc(x_281); +lean_dec(x_280); lean_inc(x_1); -x_23 = l_Lean_MessageData_ofExpr(x_1); -x_24 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -lean_ctor_set_tag(x_14, 7); -lean_ctor_set(x_14, 1, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Meta_Grind_internalize___lambda__2(x_1, x_2, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); -lean_dec(x_27); -return x_29; +x_282 = l_Lean_Meta_Grind_registerParent(x_1, x_279, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_281); +lean_dec(x_279); +x_283 = lean_ctor_get(x_282, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +lean_dec(x_282); +x_285 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_283, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_284); +lean_dec(x_283); +return x_285; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); -lean_dec(x_14); -lean_inc(x_1); -x_31 = l_Lean_MessageData_ofExpr(x_1); -x_32 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_Meta_Grind_internalize___lambda__2(x_1, x_2, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); -lean_dec(x_36); -return x_38; +uint8_t x_286; +lean_dec(x_279); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_286 = !lean_is_exclusive(x_280); +if (x_286 == 0) +{ +return x_280; } +else +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; +x_287 = lean_ctor_get(x_280, 0); +x_288 = lean_ctor_get(x_280, 1); +lean_inc(x_288); +lean_inc(x_287); +lean_dec(x_280); +x_289 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_289, 0, x_287); +lean_ctor_set(x_289, 1, x_288); +return x_289; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = l_Lean_Meta_Grind_alreadyInternalized(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); +lean_object* x_290; lean_object* x_291; +x_290 = lean_array_fget(x_4, x_276); +lean_dec(x_4); lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 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_6); +lean_inc(x_2); +lean_inc(x_290); +x_291 = l_Lean_Meta_Grind_internalize(x_290, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_257); +if (lean_obj_tag(x_291) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_box(0); -x_17 = l_Lean_Meta_Grind_internalize___lambda__3(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_17; +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +x_292 = lean_ctor_get(x_291, 1); +lean_inc(x_292); +lean_dec(x_291); +lean_inc(x_1); +x_293 = l_Lean_Meta_Grind_registerParent(x_1, x_290, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_292); +lean_dec(x_290); +x_294 = lean_ctor_get(x_293, 0); +lean_inc(x_294); +x_295 = lean_ctor_get(x_293, 1); +lean_inc(x_295); +lean_dec(x_293); +x_296 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_294, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_295); +lean_dec(x_294); +return x_296; } else { -uint8_t x_18; +uint8_t x_297; +lean_dec(x_290); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) +x_297 = !lean_is_exclusive(x_291); +if (x_297 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_12, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_12, 0, x_20); -return x_12; +return x_291; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; +lean_object* x_298; lean_object* x_299; lean_object* x_300; +x_298 = lean_ctor_get(x_291, 0); +x_299 = lean_ctor_get(x_291, 1); +lean_inc(x_299); +lean_inc(x_298); +lean_dec(x_291); +x_300 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_300, 0, x_298); +lean_ctor_set(x_300, 1, x_299); +return x_300; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2; -x_7 = lean_usize_land(x_2, x_6); -x_8 = lean_usize_to_nat(x_7); -x_9 = lean_box(2); -x_10 = lean_array_get(x_9, x_4, x_8); -switch (lean_obj_tag(x_10)) { -case 0: +} +block_268: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); +lean_object* x_260; +lean_dec(x_259); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_name_eq(x_3, x_11); -lean_dec(x_11); -if (x_12 == 0) +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +x_260 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(x_251, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_257); +lean_dec(x_251); +if (lean_obj_tag(x_260) == 0) { -lean_dec(x_8); +lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_261 = lean_ctor_get(x_260, 0); +lean_inc(x_261); +x_262 = lean_ctor_get(x_260, 1); +lean_inc(x_262); +lean_dec(x_260); +x_263 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_258, x_261, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_262); +lean_dec(x_261); lean_dec(x_4); -return x_1; +return x_263; } else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_1); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_1, 0); -lean_dec(x_14); -x_15 = lean_array_set(x_4, x_8, x_9); +uint8_t x_264; +lean_dec(x_258); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_15); -return x_1; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_264 = !lean_is_exclusive(x_260); +if (x_264 == 0) +{ +return x_260; } else { -lean_object* x_16; lean_object* x_17; -lean_dec(x_1); -x_16 = lean_array_set(x_4, x_8, x_9); -lean_dec(x_8); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; +lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_265 = lean_ctor_get(x_260, 0); +x_266 = lean_ctor_get(x_260, 1); +lean_inc(x_266); +lean_inc(x_265); +lean_dec(x_260); +x_267 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_267, 0, x_265); +lean_ctor_set(x_267, 1, x_266); +return x_267; } } } -case 1: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_1); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_1, 0); -lean_dec(x_19); -x_20 = !lean_is_exclusive(x_10); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_10, 0); -x_22 = lean_array_set(x_4, x_8, x_9); -x_23 = lean_usize_shift_right(x_2, x_5); -x_24 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_21, x_23, x_3); -lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; -lean_ctor_set(x_10, 0, x_24); -x_26 = lean_array_set(x_22, x_8, x_10); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_26); -return x_1; } else { -lean_object* x_27; uint8_t x_28; -lean_dec(x_24); -lean_free_object(x_10); -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; -x_29 = lean_array_set(x_22, x_8, x_27); +uint8_t x_301; +lean_dec(x_251); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_29); -return x_1; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_301 = !lean_is_exclusive(x_256); +if (x_301 == 0) +{ +return x_256; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_27, 0); -x_31 = lean_ctor_get(x_27, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_27); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_array_set(x_22, x_8, x_32); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_33); -return x_1; +lean_object* x_302; lean_object* x_303; lean_object* x_304; +x_302 = lean_ctor_get(x_256, 0); +x_303 = lean_ctor_get(x_256, 1); +lean_inc(x_303); +lean_inc(x_302); +lean_dec(x_256); +x_304 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_304, 0, x_302); +lean_ctor_set(x_304, 1, x_303); +return x_304; } } } else { -lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -x_34 = lean_ctor_get(x_10, 0); -lean_inc(x_34); +uint8_t x_305; +lean_dec(x_251); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); -x_35 = lean_array_set(x_4, x_8, x_9); -x_36 = lean_usize_shift_right(x_2, x_5); -x_37 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_34, x_36, x_3); -lean_inc(x_37); -x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_37); -x_40 = lean_array_set(x_35, x_8, x_39); +lean_dec(x_9); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_40); -return x_1; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -lean_dec(x_37); -x_41 = lean_ctor_get(x_38, 0); -lean_inc(x_41); -lean_dec(x_38); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - x_44 = x_41; -} else { - lean_dec_ref(x_41); - x_44 = lean_box(0); -} -if (lean_is_scalar(x_44)) { - x_45 = lean_alloc_ctor(0, 2, 0); -} else { - x_45 = x_44; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_305 = !lean_is_exclusive(x_252); +if (x_305 == 0) +{ +return x_252; } -lean_ctor_set(x_45, 0, x_42); -lean_ctor_set(x_45, 1, x_43); -x_46 = lean_array_set(x_35, x_8, x_45); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_46); -return x_1; +else +{ +lean_object* x_306; lean_object* x_307; lean_object* x_308; +x_306 = lean_ctor_get(x_252, 0); +x_307 = lean_ctor_get(x_252, 1); +lean_inc(x_307); +lean_inc(x_306); +lean_dec(x_252); +x_308 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_308, 0, x_306); +lean_ctor_set(x_308, 1, x_307); +return x_308; } } } -else +case 5: { -lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; -lean_dec(x_1); -x_47 = lean_ctor_get(x_10, 0); -lean_inc(x_47); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - x_48 = x_10; -} else { - lean_dec_ref(x_10); - x_48 = lean_box(0); +lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; +x_309 = lean_ctor_get(x_3, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_3, 1); +lean_inc(x_310); +lean_dec(x_3); +x_311 = lean_array_set(x_4, x_5, x_310); +x_312 = lean_unsigned_to_nat(1u); +x_313 = lean_nat_sub(x_5, x_312); +lean_dec(x_5); +x_3 = x_309; +x_4 = x_311; +x_5 = x_313; +goto _start; } -x_49 = lean_array_set(x_4, x_8, x_9); -x_50 = lean_usize_shift_right(x_2, x_5); -x_51 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_47, x_50, x_3); -lean_inc(x_51); -x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); -if (lean_obj_tag(x_52) == 0) +case 6: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -if (lean_is_scalar(x_48)) { - x_53 = lean_alloc_ctor(1, 1, 0); -} else { - x_53 = x_48; -} -lean_ctor_set(x_53, 0, x_51); -x_54 = lean_array_set(x_49, x_8, x_53); -lean_dec(x_8); -x_55 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_55, 0, x_54); -return x_55; +lean_object* x_315; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_315 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_316 = lean_ctor_get(x_315, 1); +lean_inc(x_316); +lean_dec(x_315); +lean_inc(x_1); +x_317 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_316); +x_318 = lean_ctor_get(x_317, 1); +lean_inc(x_318); +lean_dec(x_317); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_319 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_318); +if (lean_obj_tag(x_319) == 0) +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_334; uint8_t x_335; +x_320 = lean_ctor_get(x_319, 1); +lean_inc(x_320); +lean_dec(x_319); +lean_inc(x_2); +lean_inc(x_1); +x_321 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_321, 0, x_1); +lean_closure_set(x_321, 1, x_2); +x_334 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_335 = l_Lean_Expr_isConstOf(x_3, x_334); +if (x_335 == 0) +{ +lean_object* x_336; +x_336 = lean_box(0); +x_322 = x_336; +goto block_333; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_51); -lean_dec(x_48); -x_56 = lean_ctor_get(x_52, 0); -lean_inc(x_56); -lean_dec(x_52); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_59 = x_56; -} else { - lean_dec_ref(x_56); - x_59 = lean_box(0); +lean_object* x_337; lean_object* x_338; uint8_t x_339; +x_337 = lean_array_get_size(x_4); +x_338 = lean_unsigned_to_nat(2u); +x_339 = lean_nat_dec_eq(x_337, x_338); +if (x_339 == 0) +{ +lean_object* x_340; +lean_dec(x_337); +x_340 = lean_box(0); +x_322 = x_340; +goto block_333; } -if (lean_is_scalar(x_59)) { - x_60 = lean_alloc_ctor(0, 2, 0); -} else { - x_60 = x_59; +else +{ +lean_object* x_341; uint8_t x_342; +lean_dec(x_321); +lean_dec(x_3); +x_341 = lean_unsigned_to_nat(0u); +x_342 = lean_nat_dec_lt(x_341, x_337); +lean_dec(x_337); +if (x_342 == 0) +{ +lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_4); +x_343 = l_Lean_instInhabitedExpr; +x_344 = l_outOfBounds___rarg(x_343); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_344); +x_345 = l_Lean_Meta_Grind_internalize(x_344, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_320); +if (lean_obj_tag(x_345) == 0) +{ +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; +x_346 = lean_ctor_get(x_345, 1); +lean_inc(x_346); +lean_dec(x_345); +lean_inc(x_1); +x_347 = l_Lean_Meta_Grind_registerParent(x_1, x_344, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_346); +lean_dec(x_344); +x_348 = lean_ctor_get(x_347, 0); +lean_inc(x_348); +x_349 = lean_ctor_get(x_347, 1); +lean_inc(x_349); +lean_dec(x_347); +x_350 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_348, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_349); +lean_dec(x_348); +return x_350; } -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_58); -x_61 = lean_array_set(x_49, x_8, x_60); +else +{ +uint8_t x_351; +lean_dec(x_344); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -x_62 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_62, 0, x_61); -return x_62; -} -} +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_351 = !lean_is_exclusive(x_345); +if (x_351 == 0) +{ +return x_345; } -default: +else { -lean_dec(x_8); -lean_dec(x_4); -return x_1; +lean_object* x_352; lean_object* x_353; lean_object* x_354; +x_352 = lean_ctor_get(x_345, 0); +x_353 = lean_ctor_get(x_345, 1); +lean_inc(x_353); +lean_inc(x_352); +lean_dec(x_345); +x_354 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_354, 0, x_352); +lean_ctor_set(x_354, 1, x_353); +return x_354; } } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = lean_ctor_get(x_1, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_1, 1); -lean_inc(x_64); -x_65 = lean_unsigned_to_nat(0u); -x_66 = l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(x_63, x_3, x_65); -if (lean_obj_tag(x_66) == 0) -{ -lean_dec(x_64); -lean_dec(x_63); -return x_1; +lean_object* x_355; lean_object* x_356; +x_355 = lean_array_fget(x_4, x_341); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_355); +x_356 = l_Lean_Meta_Grind_internalize(x_355, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_320); +if (lean_obj_tag(x_356) == 0) +{ +lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; +x_357 = lean_ctor_get(x_356, 1); +lean_inc(x_357); +lean_dec(x_356); +lean_inc(x_1); +x_358 = l_Lean_Meta_Grind_registerParent(x_1, x_355, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_357); +lean_dec(x_355); +x_359 = lean_ctor_get(x_358, 0); +lean_inc(x_359); +x_360 = lean_ctor_get(x_358, 1); +lean_inc(x_360); +lean_dec(x_358); +x_361 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_359, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_360); +lean_dec(x_359); +return x_361; } else { -uint8_t x_67; -x_67 = !lean_is_exclusive(x_1); -if (x_67 == 0) +uint8_t x_362; +lean_dec(x_355); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_362 = !lean_is_exclusive(x_356); +if (x_362 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_68 = lean_ctor_get(x_1, 1); -lean_dec(x_68); -x_69 = lean_ctor_get(x_1, 0); -lean_dec(x_69); -x_70 = lean_ctor_get(x_66, 0); -lean_inc(x_70); -lean_dec(x_66); -lean_inc(x_70); -x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); -x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); -lean_ctor_set(x_1, 1, x_72); -lean_ctor_set(x_1, 0, x_71); -return x_1; +return x_356; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -lean_dec(x_1); -x_73 = lean_ctor_get(x_66, 0); -lean_inc(x_73); -lean_dec(x_66); -lean_inc(x_73); -x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); -x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; -} -} +lean_object* x_363; lean_object* x_364; lean_object* x_365; +x_363 = lean_ctor_get(x_356, 0); +x_364 = lean_ctor_get(x_356, 1); +lean_inc(x_364); +lean_inc(x_363); +lean_dec(x_356); +x_365 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_365, 0, x_363); +lean_ctor_set(x_365, 1, x_364); +return x_365; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Name_hash___override(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_2); -return x_5; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +block_333: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); -if (x_7 == 0) +lean_object* x_323; +lean_dec(x_322); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_323 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_320); +if (lean_obj_tag(x_323) == 0) { -uint8_t x_8; +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +x_324 = lean_ctor_get(x_323, 1); +lean_inc(x_324); +lean_dec(x_323); +lean_inc(x_1); +x_325 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_324); +lean_dec(x_3); +x_326 = lean_ctor_get(x_325, 0); +lean_inc(x_326); +x_327 = lean_ctor_get(x_325, 1); +lean_inc(x_327); +lean_dec(x_325); +x_328 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_321, x_326, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_327); +lean_dec(x_326); lean_dec(x_4); -x_8 = 0; -return x_8; +return x_328; } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_5, x_9); +uint8_t x_329; +lean_dec(x_321); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_12; -goto _start; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_329 = !lean_is_exclusive(x_323); +if (x_329 == 0) +{ +return x_323; } else { -uint8_t x_14; -lean_dec(x_4); -x_14 = 1; -return x_14; +lean_object* x_330; lean_object* x_331; lean_object* x_332; +x_330 = lean_ctor_get(x_323, 0); +x_331 = lean_ctor_get(x_323, 1); +lean_inc(x_331); +lean_inc(x_330); +lean_dec(x_323); +x_332 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_332, 0, x_330); +lean_ctor_set(x_332, 1, x_331); +return x_332; } } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) +else { -lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2; -x_7 = lean_usize_land(x_2, x_6); -x_8 = lean_usize_to_nat(x_7); -x_9 = lean_box(2); -x_10 = lean_array_get(x_9, x_4, x_8); +uint8_t x_366; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_4); -switch (lean_obj_tag(x_10)) { -case 0: +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_366 = !lean_is_exclusive(x_319); +if (x_366 == 0) { -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_3, x_11); -lean_dec(x_11); -return x_12; +return x_319; } -case 1: +else { -lean_object* x_13; size_t x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_usize_shift_right(x_2, x_5); -x_1 = x_13; -x_2 = x_14; -goto _start; -} -default: -{ -uint8_t x_16; -x_16 = 0; -return x_16; +lean_object* x_367; lean_object* x_368; lean_object* x_369; +x_367 = lean_ctor_get(x_319, 0); +x_368 = lean_ctor_get(x_319, 1); +lean_inc(x_368); +lean_inc(x_367); +lean_dec(x_319); +x_369 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_369, 0, x_367); +lean_ctor_set(x_369, 1, x_368); +return x_369; } } } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 1); -lean_inc(x_18); +uint8_t x_370; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_17, x_18, lean_box(0), x_19, x_3); -lean_dec(x_18); -lean_dec(x_17); -return x_20; -} -} -} -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint64_t x_3; size_t x_4; uint8_t x_5; -x_3 = l_Lean_HeadIndex_hash(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_1, x_4, x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_2) == 0) +x_370 = !lean_is_exclusive(x_315); +if (x_370 == 0) { -lean_object* x_4; -lean_dec(x_1); -x_4 = l_List_reverse___rarg(x_3); -return x_4; +return x_315; } else { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_2); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_2, 0); -x_7 = lean_ctor_get(x_2, 1); -lean_inc(x_1); -x_8 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_6); -if (x_8 == 0) -{ -lean_ctor_set(x_2, 1, x_3); -{ -lean_object* _tmp_1 = x_7; -lean_object* _tmp_2 = x_2; -x_2 = _tmp_1; -x_3 = _tmp_2; -} -goto _start; +lean_object* x_371; lean_object* x_372; lean_object* x_373; +x_371 = lean_ctor_get(x_315, 0); +x_372 = lean_ctor_get(x_315, 1); +lean_inc(x_372); +lean_inc(x_371); +lean_dec(x_315); +x_373 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_373, 0, x_371); +lean_ctor_set(x_373, 1, x_372); +return x_373; } -else -{ -lean_free_object(x_2); -lean_dec(x_6); -x_2 = x_7; -goto _start; } } -else +case 7: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_2, 0); -x_12 = lean_ctor_get(x_2, 1); +lean_object* x_374; +lean_dec(x_5); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_dec(x_2); +lean_inc(x_10); lean_inc(x_1); -x_13 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_11); -if (x_13 == 0) +x_374 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_374) == 0) { -lean_object* x_14; -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_11); -lean_ctor_set(x_14, 1, x_3); -x_2 = x_12; -x_3 = x_14; -goto _start; -} -else +lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; +x_375 = lean_ctor_get(x_374, 1); +lean_inc(x_375); +lean_dec(x_374); +lean_inc(x_1); +x_376 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_375); +x_377 = lean_ctor_get(x_376, 1); +lean_inc(x_377); +lean_dec(x_376); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_378 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_377); +if (lean_obj_tag(x_378) == 0) { -lean_dec(x_11); -x_2 = x_12; -goto _start; -} -} -} -} +lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_393; uint8_t x_394; +x_379 = lean_ctor_get(x_378, 1); +lean_inc(x_379); +lean_dec(x_378); +lean_inc(x_2); +lean_inc(x_1); +x_380 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_380, 0, x_1); +lean_closure_set(x_380, 1, x_2); +x_393 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_394 = l_Lean_Expr_isConstOf(x_3, x_393); +if (x_394 == 0) +{ +lean_object* x_395; +x_395 = lean_box(0); +x_381 = x_395; +goto block_392; } -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -if (lean_obj_tag(x_2) == 0) +lean_object* x_396; lean_object* x_397; uint8_t x_398; +x_396 = lean_array_get_size(x_4); +x_397 = lean_unsigned_to_nat(2u); +x_398 = lean_nat_dec_eq(x_396, x_397); +if (x_398 == 0) { -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_13 = l_List_reverse___rarg(x_3); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; +lean_object* x_399; +lean_dec(x_396); +x_399 = lean_box(0); +x_381 = x_399; +goto block_392; } else { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_2); -if (x_15 == 0) +lean_object* x_400; uint8_t x_401; +lean_dec(x_380); +lean_dec(x_3); +x_400 = lean_unsigned_to_nat(0u); +x_401 = lean_nat_dec_lt(x_400, x_396); +lean_dec(x_396); +if (x_401 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_2, 0); -x_17 = lean_ctor_get(x_2, 1); +lean_object* x_402; lean_object* x_403; lean_object* x_404; +lean_dec(x_4); +x_402 = l_Lean_instInhabitedExpr; +x_403 = l_outOfBounds___rarg(x_402); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_16, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -lean_ctor_set(x_2, 1, x_3); -lean_ctor_set(x_2, 0, x_19); +lean_inc(x_2); +lean_inc(x_403); +x_404 = l_Lean_Meta_Grind_internalize(x_403, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_379); +if (lean_obj_tag(x_404) == 0) { -lean_object* _tmp_1 = x_17; -lean_object* _tmp_2 = x_2; -lean_object* _tmp_11 = x_20; -x_2 = _tmp_1; -x_3 = _tmp_2; -x_12 = _tmp_11; -} -goto _start; +lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; +x_405 = lean_ctor_get(x_404, 1); +lean_inc(x_405); +lean_dec(x_404); +lean_inc(x_1); +x_406 = l_Lean_Meta_Grind_registerParent(x_1, x_403, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_405); +lean_dec(x_403); +x_407 = lean_ctor_get(x_406, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_406, 1); +lean_inc(x_408); +lean_dec(x_406); +x_409 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_407, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_408); +lean_dec(x_407); +return x_409; } else { -uint8_t x_22; -lean_free_object(x_2); -lean_dec(x_17); +uint8_t x_410; +lean_dec(x_403); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_22 = !lean_is_exclusive(x_18); -if (x_22 == 0) +x_410 = !lean_is_exclusive(x_404); +if (x_410 == 0) { -return x_18; +return x_404; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_18, 0); -x_24 = lean_ctor_get(x_18, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_18); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_411; lean_object* x_412; lean_object* x_413; +x_411 = lean_ctor_get(x_404, 0); +x_412 = lean_ctor_get(x_404, 1); +lean_inc(x_412); +lean_inc(x_411); +lean_dec(x_404); +x_413 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_413, 0, x_411); +lean_ctor_set(x_413, 1, x_412); +return x_413; } } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_2, 0); -x_27 = lean_ctor_get(x_2, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_2); +lean_object* x_414; lean_object* x_415; +x_414 = lean_array_fget(x_4, x_400); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_1); -x_28 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_26, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_28) == 0) +lean_inc(x_2); +lean_inc(x_414); +x_415 = l_Lean_Meta_Grind_internalize(x_414, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_379); +if (lean_obj_tag(x_415) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_3); -x_2 = x_27; -x_3 = x_31; -x_12 = x_30; -goto _start; +lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; +x_416 = lean_ctor_get(x_415, 1); +lean_inc(x_416); +lean_dec(x_415); +lean_inc(x_1); +x_417 = l_Lean_Meta_Grind_registerParent(x_1, x_414, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_416); +lean_dec(x_414); +x_418 = lean_ctor_get(x_417, 0); +lean_inc(x_418); +x_419 = lean_ctor_get(x_417, 1); +lean_inc(x_419); +lean_dec(x_417); +x_420 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_418, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_419); +lean_dec(x_418); +return x_420; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_27); +uint8_t x_421; +lean_dec(x_414); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_33 = lean_ctor_get(x_28, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_28, 1); -lean_inc(x_34); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_35 = x_28; -} else { - lean_dec_ref(x_28); - x_35 = lean_box(0); -} -if (lean_is_scalar(x_35)) { - x_36 = lean_alloc_ctor(1, 2, 0); -} else { - x_36 = x_35; +x_421 = !lean_is_exclusive(x_415); +if (x_421 == 0) +{ +return x_415; } -lean_ctor_set(x_36, 0, x_33); -lean_ctor_set(x_36, 1, x_34); -return x_36; +else +{ +lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_422 = lean_ctor_get(x_415, 0); +x_423 = lean_ctor_get(x_415, 1); +lean_inc(x_423); +lean_inc(x_422); +lean_dec(x_415); +x_424 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_424, 0, x_422); +lean_ctor_set(x_424, 1, x_423); +return x_424; } } } } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1() { -_start: +block_392: { -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_382; +lean_dec(x_381); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_382 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_379); +if (lean_obj_tag(x_382) == 0) +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +x_383 = lean_ctor_get(x_382, 1); +lean_inc(x_383); +lean_dec(x_382); +lean_inc(x_1); +x_384 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_383); +lean_dec(x_3); +x_385 = lean_ctor_get(x_384, 0); +lean_inc(x_385); +x_386 = lean_ctor_get(x_384, 1); +lean_inc(x_386); +lean_dec(x_384); +x_387 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_380, x_385, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_386); +lean_dec(x_385); +lean_dec(x_4); +return x_387; } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_st_ref_take(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); +uint8_t x_388; +lean_dec(x_380); +lean_dec(x_13); lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_13, 9); -x_17 = l_Lean_PersistentArray_push___rarg(x_16, x_1); -lean_ctor_set(x_13, 9, x_17); -x_18 = lean_st_ref_set(x_3, x_13, x_14); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_388 = !lean_is_exclusive(x_382); +if (x_388 == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -lean_ctor_set(x_18, 0, x_21); -return x_18; +return x_382; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; +lean_object* x_389; lean_object* x_390; lean_object* x_391; +x_389 = lean_ctor_get(x_382, 0); +x_390 = lean_ctor_get(x_382, 1); +lean_inc(x_390); +lean_inc(x_389); +lean_dec(x_382); +x_391 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_391, 0, x_389); +lean_ctor_set(x_391, 1, x_390); +return x_391; +} +} } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_25 = lean_ctor_get(x_13, 0); -x_26 = lean_ctor_get(x_13, 1); -x_27 = lean_ctor_get(x_13, 2); -x_28 = lean_ctor_get(x_13, 3); -x_29 = lean_ctor_get(x_13, 4); -x_30 = lean_ctor_get(x_13, 5); -x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); -x_32 = lean_ctor_get(x_13, 6); -x_33 = lean_ctor_get(x_13, 7); -x_34 = lean_ctor_get(x_13, 8); -x_35 = lean_ctor_get(x_13, 9); -x_36 = lean_ctor_get(x_13, 10); -x_37 = lean_ctor_get(x_13, 11); -x_38 = lean_ctor_get(x_13, 12); -x_39 = lean_ctor_get(x_13, 13); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); +uint8_t x_425; lean_dec(x_13); -x_40 = l_Lean_PersistentArray_push___rarg(x_35, x_1); -x_41 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_26); -lean_ctor_set(x_41, 2, x_27); -lean_ctor_set(x_41, 3, x_28); -lean_ctor_set(x_41, 4, x_29); -lean_ctor_set(x_41, 5, x_30); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_33); -lean_ctor_set(x_41, 8, x_34); -lean_ctor_set(x_41, 9, x_40); -lean_ctor_set(x_41, 10, x_36); -lean_ctor_set(x_41, 11, x_37); -lean_ctor_set(x_41, 12, x_38); -lean_ctor_set(x_41, 13, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); -x_42 = lean_st_ref_set(x_3, x_41, x_14); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); -} -x_45 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_425 = !lean_is_exclusive(x_378); +if (x_425 == 0) +{ +return x_378; } -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +else +{ +lean_object* x_426; lean_object* x_427; lean_object* x_428; +x_426 = lean_ctor_get(x_378, 0); +x_427 = lean_ctor_get(x_378, 1); +lean_inc(x_427); +lean_inc(x_426); +lean_dec(x_378); +x_428 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_428, 0, x_426); +lean_ctor_set(x_428, 1, x_427); +return x_428; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_st_ref_take(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); +uint8_t x_429; +lean_dec(x_13); lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_429 = !lean_is_exclusive(x_374); +if (x_429 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_13, 10); -x_17 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_16, x_1); -lean_ctor_set(x_13, 10, x_17); -x_18 = lean_st_ref_set(x_3, x_13, x_14); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +return x_374; +} +else { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -lean_ctor_set(x_18, 0, x_21); -return x_18; +lean_object* x_430; lean_object* x_431; lean_object* x_432; +x_430 = lean_ctor_get(x_374, 0); +x_431 = lean_ctor_get(x_374, 1); +lean_inc(x_431); +lean_inc(x_430); +lean_dec(x_374); +x_432 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_432, 0, x_430); +lean_ctor_set(x_432, 1, x_431); +return x_432; +} +} +} +case 8: +{ +lean_object* x_433; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_433 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_433) == 0) +{ +lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; +x_434 = lean_ctor_get(x_433, 1); +lean_inc(x_434); +lean_dec(x_433); +lean_inc(x_1); +x_435 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_434); +x_436 = lean_ctor_get(x_435, 1); +lean_inc(x_436); +lean_dec(x_435); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_437 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_436); +if (lean_obj_tag(x_437) == 0) +{ +lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_452; uint8_t x_453; +x_438 = lean_ctor_get(x_437, 1); +lean_inc(x_438); +lean_dec(x_437); +lean_inc(x_2); +lean_inc(x_1); +x_439 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_439, 0, x_1); +lean_closure_set(x_439, 1, x_2); +x_452 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_453 = l_Lean_Expr_isConstOf(x_3, x_452); +if (x_453 == 0) +{ +lean_object* x_454; +x_454 = lean_box(0); +x_440 = x_454; +goto block_451; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; +lean_object* x_455; lean_object* x_456; uint8_t x_457; +x_455 = lean_array_get_size(x_4); +x_456 = lean_unsigned_to_nat(2u); +x_457 = lean_nat_dec_eq(x_455, x_456); +if (x_457 == 0) +{ +lean_object* x_458; +lean_dec(x_455); +x_458 = lean_box(0); +x_440 = x_458; +goto block_451; } +else +{ +lean_object* x_459; uint8_t x_460; +lean_dec(x_439); +lean_dec(x_3); +x_459 = lean_unsigned_to_nat(0u); +x_460 = lean_nat_dec_lt(x_459, x_455); +lean_dec(x_455); +if (x_460 == 0) +{ +lean_object* x_461; lean_object* x_462; lean_object* x_463; +lean_dec(x_4); +x_461 = l_Lean_instInhabitedExpr; +x_462 = l_outOfBounds___rarg(x_461); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_462); +x_463 = l_Lean_Meta_Grind_internalize(x_462, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_438); +if (lean_obj_tag(x_463) == 0) +{ +lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; +x_464 = lean_ctor_get(x_463, 1); +lean_inc(x_464); +lean_dec(x_463); +lean_inc(x_1); +x_465 = l_Lean_Meta_Grind_registerParent(x_1, x_462, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_464); +lean_dec(x_462); +x_466 = lean_ctor_get(x_465, 0); +lean_inc(x_466); +x_467 = lean_ctor_get(x_465, 1); +lean_inc(x_467); +lean_dec(x_465); +x_468 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_466, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_467); +lean_dec(x_466); +return x_468; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_25 = lean_ctor_get(x_13, 0); -x_26 = lean_ctor_get(x_13, 1); -x_27 = lean_ctor_get(x_13, 2); -x_28 = lean_ctor_get(x_13, 3); -x_29 = lean_ctor_get(x_13, 4); -x_30 = lean_ctor_get(x_13, 5); -x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); -x_32 = lean_ctor_get(x_13, 6); -x_33 = lean_ctor_get(x_13, 7); -x_34 = lean_ctor_get(x_13, 8); -x_35 = lean_ctor_get(x_13, 9); -x_36 = lean_ctor_get(x_13, 10); -x_37 = lean_ctor_get(x_13, 11); -x_38 = lean_ctor_get(x_13, 12); -x_39 = lean_ctor_get(x_13, 13); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); +uint8_t x_469; +lean_dec(x_462); lean_dec(x_13); -x_40 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_36, x_1); -x_41 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_26); -lean_ctor_set(x_41, 2, x_27); -lean_ctor_set(x_41, 3, x_28); -lean_ctor_set(x_41, 4, x_29); -lean_ctor_set(x_41, 5, x_30); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_33); -lean_ctor_set(x_41, 8, x_34); -lean_ctor_set(x_41, 9, x_35); -lean_ctor_set(x_41, 10, x_40); -lean_ctor_set(x_41, 11, x_37); -lean_ctor_set(x_41, 12, x_38); -lean_ctor_set(x_41, 13, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); -x_42 = lean_st_ref_set(x_3, x_41, x_14); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); -} -x_45 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1; -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_469 = !lean_is_exclusive(x_463); +if (x_469 == 0) +{ +return x_463; } -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +else +{ +lean_object* x_470; lean_object* x_471; lean_object* x_472; +x_470 = lean_ctor_get(x_463, 0); +x_471 = lean_ctor_get(x_463, 1); +lean_inc(x_471); +lean_inc(x_470); +lean_dec(x_463); +x_472 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_472, 0, x_470); +lean_ctor_set(x_472, 1, x_471); +return x_472; } } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ematch", 6, 6); -return x_1; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2() { -_start: +lean_object* x_473; lean_object* x_474; +x_473 = lean_array_fget(x_4, x_459); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_473); +x_474 = l_Lean_Meta_Grind_internalize(x_473, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_438); +if (lean_obj_tag(x_474) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; -x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} +lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; +x_475 = lean_ctor_get(x_474, 1); +lean_inc(x_475); +lean_dec(x_474); +lean_inc(x_1); +x_476 = l_Lean_Meta_Grind_registerParent(x_1, x_473, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_475); +lean_dec(x_473); +x_477 = lean_ctor_get(x_476, 0); +lean_inc(x_477); +x_478 = lean_ctor_get(x_476, 1); +lean_inc(x_478); +lean_dec(x_476); +x_479 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_477, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_478); +lean_dec(x_477); +return x_479; } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("activated `", 11, 11); -return x_1; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4() { -_start: +uint8_t x_480; +lean_dec(x_473); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_480 = !lean_is_exclusive(x_474); +if (x_480 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +return x_474; } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`, ", 3, 3); -return x_1; -} +lean_object* x_481; lean_object* x_482; lean_object* x_483; +x_481 = lean_ctor_get(x_474, 0); +x_482 = lean_ctor_get(x_474, 1); +lean_inc(x_482); +lean_inc(x_481); +lean_dec(x_474); +x_483 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_483, 0, x_481); +lean_ctor_set(x_483, 1, x_482); +return x_483; } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("reinsert `", 10, 10); -return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8() { -_start: +block_451: { -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_441; +lean_dec(x_440); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_441 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_438); +if (lean_obj_tag(x_441) == 0) +{ +lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; +x_442 = lean_ctor_get(x_441, 1); +lean_inc(x_442); +lean_dec(x_441); +lean_inc(x_1); +x_443 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_442); +lean_dec(x_3); +x_444 = lean_ctor_get(x_443, 0); +lean_inc(x_444); +x_445 = lean_ctor_get(x_443, 1); +lean_inc(x_445); +lean_dec(x_443); +x_446 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_439, x_444, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_445); +lean_dec(x_444); +lean_dec(x_4); +return x_446; } +else +{ +uint8_t x_447; +lean_dec(x_439); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_447 = !lean_is_exclusive(x_441); +if (x_447 == 0) +{ +return x_441; } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`", 1, 1); -return x_1; +lean_object* x_448; lean_object* x_449; lean_object* x_450; +x_448 = lean_ctor_get(x_441, 0); +x_449 = lean_ctor_get(x_441, 1); +lean_inc(x_449); +lean_inc(x_448); +lean_dec(x_441); +x_450 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_450, 0, x_448); +lean_ctor_set(x_450, 1, x_449); +return x_450; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: -{ -if (lean_obj_tag(x_6) == 0) +else { -lean_object* x_18; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +uint8_t x_484; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_17); -return x_18; +x_484 = !lean_is_exclusive(x_437); +if (x_484 == 0) +{ +return x_437; } else { -uint8_t x_19; -lean_dec(x_7); -x_19 = !lean_is_exclusive(x_6); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_27; -x_20 = lean_ctor_get(x_6, 0); -x_21 = lean_ctor_get(x_6, 1); -x_27 = !lean_is_exclusive(x_20); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_28 = lean_ctor_get(x_20, 0); -x_29 = lean_ctor_get(x_20, 1); -x_30 = lean_ctor_get(x_20, 2); -x_31 = lean_ctor_get(x_20, 3); -x_32 = lean_ctor_get(x_20, 4); -x_33 = lean_ctor_get(x_20, 5); -x_34 = lean_box(0); -lean_inc(x_3); -x_35 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_32, x_34); -if (lean_obj_tag(x_35) == 0) +lean_object* x_485; lean_object* x_486; lean_object* x_487; +x_485 = lean_ctor_get(x_437, 0); +x_486 = lean_ctor_get(x_437, 1); +lean_inc(x_486); +lean_inc(x_485); +lean_dec(x_437); +x_487 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_487, 0, x_485); +lean_ctor_set(x_487, 1, x_486); +return x_487; +} +} +} +else { -lean_object* x_36; uint8_t x_37; -x_36 = l_Lean_Meta_Grind_shareCommon(x_29, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) +uint8_t x_488; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_488 = !lean_is_exclusive(x_433); +if (x_488 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_36, 0); -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); +return x_433; +} +else +{ +lean_object* x_489; lean_object* x_490; lean_object* x_491; +x_489 = lean_ctor_get(x_433, 0); +x_490 = lean_ctor_get(x_433, 1); +lean_inc(x_490); +lean_inc(x_489); +lean_dec(x_433); +x_491 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_491, 0, x_489); +lean_ctor_set(x_491, 1, x_490); +return x_491; +} +} +} +case 9: +{ +lean_object* x_492; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_492 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_492) == 0) +{ +lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; +x_493 = lean_ctor_get(x_492, 1); +lean_inc(x_493); +lean_dec(x_492); +lean_inc(x_1); +x_494 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_493); +x_495 = lean_ctor_get(x_494, 1); +lean_inc(x_495); +lean_dec(x_494); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_496 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_495); +if (lean_obj_tag(x_496) == 0) +{ +lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_511; uint8_t x_512; +x_497 = lean_ctor_get(x_496, 1); +lean_inc(x_497); +lean_dec(x_496); +lean_inc(x_2); lean_inc(x_1); -x_40 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_31, x_34, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_39); -if (lean_obj_tag(x_40) == 0) +x_498 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_498, 0, x_1); +lean_closure_set(x_498, 1, x_2); +x_511 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_512 = l_Lean_Expr_isConstOf(x_3, x_511); +if (x_512 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -lean_inc(x_33); -lean_inc(x_41); -lean_ctor_set(x_20, 4, x_35); -lean_ctor_set(x_20, 3, x_41); -lean_ctor_set(x_20, 1, x_38); -x_43 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_44 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_43, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_42); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_unbox(x_45); -lean_dec(x_45); -if (x_46 == 0) +lean_object* x_513; +x_513 = lean_box(0); +x_499 = x_513; +goto block_510; +} +else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_41); -lean_free_object(x_36); -lean_dec(x_33); -lean_free_object(x_6); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_box(0); -x_49 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_47); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_22 = x_50; -x_23 = x_51; -goto block_26; +lean_object* x_514; lean_object* x_515; uint8_t x_516; +x_514 = lean_array_get_size(x_4); +x_515 = lean_unsigned_to_nat(2u); +x_516 = lean_nat_dec_eq(x_514, x_515); +if (x_516 == 0) +{ +lean_object* x_517; +lean_dec(x_514); +x_517 = lean_box(0); +x_499 = x_517; +goto block_510; } else { -uint8_t x_52; -x_52 = !lean_is_exclusive(x_44); -if (x_52 == 0) +lean_object* x_518; uint8_t x_519; +lean_dec(x_498); +lean_dec(x_3); +x_518 = lean_unsigned_to_nat(0u); +x_519 = lean_nat_dec_lt(x_518, x_514); +lean_dec(x_514); +if (x_519 == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_53 = lean_ctor_get(x_44, 1); -x_54 = lean_ctor_get(x_44, 0); -lean_dec(x_54); -x_55 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_56 = l_Lean_MessageData_ofName(x_55); -x_57 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -lean_ctor_set_tag(x_44, 7); -lean_ctor_set(x_44, 1, x_56); -lean_ctor_set(x_44, 0, x_57); -x_58 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -lean_ctor_set_tag(x_36, 7); -lean_ctor_set(x_36, 1, x_58); -lean_ctor_set(x_36, 0, x_44); -x_59 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_41, x_34); -x_60 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_59, x_34); -x_61 = l_Lean_MessageData_ofList(x_60); -lean_ctor_set_tag(x_6, 7); -lean_ctor_set(x_6, 1, x_61); -lean_ctor_set(x_6, 0, x_36); -x_62 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_6); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_43, x_63, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_53); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_65, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_66); -lean_dec(x_65); -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -x_22 = x_68; -x_23 = x_69; -goto block_26; +lean_object* x_520; lean_object* x_521; lean_object* x_522; +lean_dec(x_4); +x_520 = l_Lean_instInhabitedExpr; +x_521 = l_outOfBounds___rarg(x_520); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_521); +x_522 = l_Lean_Meta_Grind_internalize(x_521, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_497); +if (lean_obj_tag(x_522) == 0) +{ +lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; +x_523 = lean_ctor_get(x_522, 1); +lean_inc(x_523); +lean_dec(x_522); +lean_inc(x_1); +x_524 = l_Lean_Meta_Grind_registerParent(x_1, x_521, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_523); +lean_dec(x_521); +x_525 = lean_ctor_get(x_524, 0); +lean_inc(x_525); +x_526 = lean_ctor_get(x_524, 1); +lean_inc(x_526); +lean_dec(x_524); +x_527 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_525, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_526); +lean_dec(x_525); +return x_527; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_70 = lean_ctor_get(x_44, 1); -lean_inc(x_70); -lean_dec(x_44); -x_71 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_72 = l_Lean_MessageData_ofName(x_71); -x_73 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -x_74 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_72); -x_75 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -lean_ctor_set_tag(x_36, 7); -lean_ctor_set(x_36, 1, x_75); -lean_ctor_set(x_36, 0, x_74); -x_76 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_41, x_34); -x_77 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_76, x_34); -x_78 = l_Lean_MessageData_ofList(x_77); -lean_ctor_set_tag(x_6, 7); -lean_ctor_set(x_6, 1, x_78); -lean_ctor_set(x_6, 0, x_36); -x_79 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_6); -lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_43, x_80, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_70); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_82, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_83); -lean_dec(x_82); -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_84, 1); -lean_inc(x_86); -lean_dec(x_84); -x_22 = x_85; -x_23 = x_86; -goto block_26; +uint8_t x_528; +lean_dec(x_521); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_528 = !lean_is_exclusive(x_522); +if (x_528 == 0) +{ +return x_522; +} +else +{ +lean_object* x_529; lean_object* x_530; lean_object* x_531; +x_529 = lean_ctor_get(x_522, 0); +x_530 = lean_ctor_get(x_522, 1); +lean_inc(x_530); +lean_inc(x_529); +lean_dec(x_522); +x_531 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_531, 0, x_529); +lean_ctor_set(x_531, 1, x_530); +return x_531; } } } else { -uint8_t x_87; -lean_free_object(x_36); -lean_dec(x_38); -lean_free_object(x_20); -lean_dec(x_33); -lean_dec(x_30); -lean_dec(x_28); -lean_free_object(x_6); -lean_dec(x_21); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +lean_object* x_532; lean_object* x_533; +x_532 = lean_array_fget(x_4, x_518); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_532); +x_533 = l_Lean_Meta_Grind_internalize(x_532, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_497); +if (lean_obj_tag(x_533) == 0) +{ +lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; +x_534 = lean_ctor_get(x_533, 1); +lean_inc(x_534); +lean_dec(x_533); +lean_inc(x_1); +x_535 = l_Lean_Meta_Grind_registerParent(x_1, x_532, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_534); +lean_dec(x_532); +x_536 = lean_ctor_get(x_535, 0); +lean_inc(x_536); +x_537 = lean_ctor_get(x_535, 1); +lean_inc(x_537); +lean_dec(x_535); +x_538 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_536, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_537); +lean_dec(x_536); +return x_538; +} +else +{ +uint8_t x_539; +lean_dec(x_532); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_3); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); lean_dec(x_1); -x_87 = !lean_is_exclusive(x_40); -if (x_87 == 0) +x_539 = !lean_is_exclusive(x_533); +if (x_539 == 0) { -return x_40; +return x_533; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_40, 0); -x_89 = lean_ctor_get(x_40, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_40); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_88); -lean_ctor_set(x_90, 1, x_89); -return x_90; +lean_object* x_540; lean_object* x_541; lean_object* x_542; +x_540 = lean_ctor_get(x_533, 0); +x_541 = lean_ctor_get(x_533, 1); +lean_inc(x_541); +lean_inc(x_540); +lean_dec(x_533); +x_542 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_542, 0, x_540); +lean_ctor_set(x_542, 1, x_541); +return x_542; } } } -else +} +} +block_510: { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_36, 0); -x_92 = lean_ctor_get(x_36, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_36); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); +lean_object* x_500; +lean_dec(x_499); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_500 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_497); +if (lean_obj_tag(x_500) == 0) +{ +lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; +x_501 = lean_ctor_get(x_500, 1); +lean_inc(x_501); +lean_dec(x_500); lean_inc(x_1); -x_93 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_31, x_34, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_92); -if (lean_obj_tag(x_93) == 0) +x_502 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_501); +lean_dec(x_3); +x_503 = lean_ctor_get(x_502, 0); +lean_inc(x_503); +x_504 = lean_ctor_get(x_502, 1); +lean_inc(x_504); +lean_dec(x_502); +x_505 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_498, x_503, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_504); +lean_dec(x_503); +lean_dec(x_4); +return x_505; +} +else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_94 = lean_ctor_get(x_93, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -lean_dec(x_93); -lean_inc(x_33); -lean_inc(x_94); -lean_ctor_set(x_20, 4, x_35); -lean_ctor_set(x_20, 3, x_94); -lean_ctor_set(x_20, 1, x_91); -x_96 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_97 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_96, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_95); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_unbox(x_98); -lean_dec(x_98); -if (x_99 == 0) +uint8_t x_506; +lean_dec(x_498); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_506 = !lean_is_exclusive(x_500); +if (x_506 == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -lean_dec(x_94); -lean_dec(x_33); -lean_free_object(x_6); -x_100 = lean_ctor_get(x_97, 1); -lean_inc(x_100); -lean_dec(x_97); -x_101 = lean_box(0); -x_102 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_101, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_100); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_22 = x_103; -x_23 = x_104; -goto block_26; +return x_500; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_105 = lean_ctor_get(x_97, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_97)) { - lean_ctor_release(x_97, 0); - lean_ctor_release(x_97, 1); - x_106 = x_97; -} else { - lean_dec_ref(x_97); - x_106 = lean_box(0); +lean_object* x_507; lean_object* x_508; lean_object* x_509; +x_507 = lean_ctor_get(x_500, 0); +x_508 = lean_ctor_get(x_500, 1); +lean_inc(x_508); +lean_inc(x_507); +lean_dec(x_500); +x_509 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_509, 0, x_507); +lean_ctor_set(x_509, 1, x_508); +return x_509; +} } -x_107 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_108 = l_Lean_MessageData_ofName(x_107); -x_109 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -if (lean_is_scalar(x_106)) { - x_110 = lean_alloc_ctor(7, 2, 0); -} else { - x_110 = x_106; - lean_ctor_set_tag(x_110, 7); -} -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_108); -x_111 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -x_112 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -x_113 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_94, x_34); -x_114 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_113, x_34); -x_115 = l_Lean_MessageData_ofList(x_114); -lean_ctor_set_tag(x_6, 7); -lean_ctor_set(x_6, 1, x_115); -lean_ctor_set(x_6, 0, x_112); -x_116 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_117 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_117, 0, x_6); -lean_ctor_set(x_117, 1, x_116); -x_118 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_96, x_117, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_105); -x_119 = lean_ctor_get(x_118, 0); -lean_inc(x_119); -x_120 = lean_ctor_get(x_118, 1); -lean_inc(x_120); -lean_dec(x_118); -x_121 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_20, x_119, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_120); -lean_dec(x_119); -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -x_123 = lean_ctor_get(x_121, 1); -lean_inc(x_123); -lean_dec(x_121); -x_22 = x_122; -x_23 = x_123; -goto block_26; } } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_91); -lean_free_object(x_20); -lean_dec(x_33); -lean_dec(x_30); -lean_dec(x_28); -lean_free_object(x_6); -lean_dec(x_21); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +uint8_t x_543; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_124 = lean_ctor_get(x_93, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_93, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_93)) { - lean_ctor_release(x_93, 0); - lean_ctor_release(x_93, 1); - x_126 = x_93; +x_543 = !lean_is_exclusive(x_496); +if (x_543 == 0) +{ +return x_496; +} +else +{ +lean_object* x_544; lean_object* x_545; lean_object* x_546; +x_544 = lean_ctor_get(x_496, 0); +x_545 = lean_ctor_get(x_496, 1); +lean_inc(x_545); +lean_inc(x_544); +lean_dec(x_496); +x_546 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_546, 0, x_544); +lean_ctor_set(x_546, 1, x_545); +return x_546; +} +} +} +else +{ +uint8_t x_547; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_547 = !lean_is_exclusive(x_492); +if (x_547 == 0) +{ +return x_492; +} +else +{ +lean_object* x_548; lean_object* x_549; lean_object* x_550; +x_548 = lean_ctor_get(x_492, 0); +x_549 = lean_ctor_get(x_492, 1); +lean_inc(x_549); +lean_inc(x_548); +lean_dec(x_492); +x_550 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_550, 0, x_548); +lean_ctor_set(x_550, 1, x_549); +return x_550; +} +} +} +case 10: +{ +lean_object* x_551; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_551 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_551) == 0) +{ +lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; +x_552 = lean_ctor_get(x_551, 1); +lean_inc(x_552); +lean_dec(x_551); +lean_inc(x_1); +x_553 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_552); +x_554 = lean_ctor_get(x_553, 1); +lean_inc(x_554); +lean_dec(x_553); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_555 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_554); +if (lean_obj_tag(x_555) == 0) +{ +lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_570; uint8_t x_571; +x_556 = lean_ctor_get(x_555, 1); +lean_inc(x_556); +lean_dec(x_555); +lean_inc(x_2); +lean_inc(x_1); +x_557 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_557, 0, x_1); +lean_closure_set(x_557, 1, x_2); +x_570 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_571 = l_Lean_Expr_isConstOf(x_3, x_570); +if (x_571 == 0) +{ +lean_object* x_572; +x_572 = lean_box(0); +x_558 = x_572; +goto block_569; +} +else +{ +lean_object* x_573; lean_object* x_574; uint8_t x_575; +x_573 = lean_array_get_size(x_4); +x_574 = lean_unsigned_to_nat(2u); +x_575 = lean_nat_dec_eq(x_573, x_574); +if (x_575 == 0) +{ +lean_object* x_576; +lean_dec(x_573); +x_576 = lean_box(0); +x_558 = x_576; +goto block_569; +} +else +{ +lean_object* x_577; uint8_t x_578; +lean_dec(x_557); +lean_dec(x_3); +x_577 = lean_unsigned_to_nat(0u); +x_578 = lean_nat_dec_lt(x_577, x_573); +lean_dec(x_573); +if (x_578 == 0) +{ +lean_object* x_579; lean_object* x_580; lean_object* x_581; +lean_dec(x_4); +x_579 = l_Lean_instInhabitedExpr; +x_580 = l_outOfBounds___rarg(x_579); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_580); +x_581 = l_Lean_Meta_Grind_internalize(x_580, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_556); +if (lean_obj_tag(x_581) == 0) +{ +lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; +x_582 = lean_ctor_get(x_581, 1); +lean_inc(x_582); +lean_dec(x_581); +lean_inc(x_1); +x_583 = l_Lean_Meta_Grind_registerParent(x_1, x_580, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_582); +lean_dec(x_580); +x_584 = lean_ctor_get(x_583, 0); +lean_inc(x_584); +x_585 = lean_ctor_get(x_583, 1); +lean_inc(x_585); +lean_dec(x_583); +x_586 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_584, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_585); +lean_dec(x_584); +return x_586; +} +else +{ +uint8_t x_587; +lean_dec(x_580); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_587 = !lean_is_exclusive(x_581); +if (x_587 == 0) +{ +return x_581; +} +else +{ +lean_object* x_588; lean_object* x_589; lean_object* x_590; +x_588 = lean_ctor_get(x_581, 0); +x_589 = lean_ctor_get(x_581, 1); +lean_inc(x_589); +lean_inc(x_588); +lean_dec(x_581); +x_590 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_590, 0, x_588); +lean_ctor_set(x_590, 1, x_589); +return x_590; +} +} +} +else +{ +lean_object* x_591; lean_object* x_592; +x_591 = lean_array_fget(x_4, x_577); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_591); +x_592 = l_Lean_Meta_Grind_internalize(x_591, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_556); +if (lean_obj_tag(x_592) == 0) +{ +lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; +x_593 = lean_ctor_get(x_592, 1); +lean_inc(x_593); +lean_dec(x_592); +lean_inc(x_1); +x_594 = l_Lean_Meta_Grind_registerParent(x_1, x_591, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_593); +lean_dec(x_591); +x_595 = lean_ctor_get(x_594, 0); +lean_inc(x_595); +x_596 = lean_ctor_get(x_594, 1); +lean_inc(x_596); +lean_dec(x_594); +x_597 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_595, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_596); +lean_dec(x_595); +return x_597; +} +else +{ +uint8_t x_598; +lean_dec(x_591); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_598 = !lean_is_exclusive(x_592); +if (x_598 == 0) +{ +return x_592; +} +else +{ +lean_object* x_599; lean_object* x_600; lean_object* x_601; +x_599 = lean_ctor_get(x_592, 0); +x_600 = lean_ctor_get(x_592, 1); +lean_inc(x_600); +lean_inc(x_599); +lean_dec(x_592); +x_601 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_601, 0, x_599); +lean_ctor_set(x_601, 1, x_600); +return x_601; +} +} +} +} +} +block_569: +{ +lean_object* x_559; +lean_dec(x_558); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_559 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_556); +if (lean_obj_tag(x_559) == 0) +{ +lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; +x_560 = lean_ctor_get(x_559, 1); +lean_inc(x_560); +lean_dec(x_559); +lean_inc(x_1); +x_561 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_560); +lean_dec(x_3); +x_562 = lean_ctor_get(x_561, 0); +lean_inc(x_562); +x_563 = lean_ctor_get(x_561, 1); +lean_inc(x_563); +lean_dec(x_561); +x_564 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_557, x_562, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_563); +lean_dec(x_562); +lean_dec(x_4); +return x_564; +} +else +{ +uint8_t x_565; +lean_dec(x_557); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_565 = !lean_is_exclusive(x_559); +if (x_565 == 0) +{ +return x_559; +} +else +{ +lean_object* x_566; lean_object* x_567; lean_object* x_568; +x_566 = lean_ctor_get(x_559, 0); +x_567 = lean_ctor_get(x_559, 1); +lean_inc(x_567); +lean_inc(x_566); +lean_dec(x_559); +x_568 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_568, 0, x_566); +lean_ctor_set(x_568, 1, x_567); +return x_568; +} +} +} +} +else +{ +uint8_t x_602; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_602 = !lean_is_exclusive(x_555); +if (x_602 == 0) +{ +return x_555; +} +else +{ +lean_object* x_603; lean_object* x_604; lean_object* x_605; +x_603 = lean_ctor_get(x_555, 0); +x_604 = lean_ctor_get(x_555, 1); +lean_inc(x_604); +lean_inc(x_603); +lean_dec(x_555); +x_605 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_605, 0, x_603); +lean_ctor_set(x_605, 1, x_604); +return x_605; +} +} +} +else +{ +uint8_t x_606; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_606 = !lean_is_exclusive(x_551); +if (x_606 == 0) +{ +return x_551; +} +else +{ +lean_object* x_607; lean_object* x_608; lean_object* x_609; +x_607 = lean_ctor_get(x_551, 0); +x_608 = lean_ctor_get(x_551, 1); +lean_inc(x_608); +lean_inc(x_607); +lean_dec(x_551); +x_609 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_609, 0, x_607); +lean_ctor_set(x_609, 1, x_608); +return x_609; +} +} +} +default: +{ +lean_object* x_610; +lean_dec(x_5); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_1); +x_610 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_610) == 0) +{ +lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; +x_611 = lean_ctor_get(x_610, 1); +lean_inc(x_611); +lean_dec(x_610); +lean_inc(x_1); +x_612 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_611); +x_613 = lean_ctor_get(x_612, 1); +lean_inc(x_613); +lean_dec(x_612); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_614 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_613); +if (lean_obj_tag(x_614) == 0) +{ +lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_629; uint8_t x_630; +x_615 = lean_ctor_get(x_614, 1); +lean_inc(x_615); +lean_dec(x_614); +lean_inc(x_2); +lean_inc(x_1); +x_616 = lean_alloc_closure((void*)(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1___boxed), 12, 2); +lean_closure_set(x_616, 0, x_1); +lean_closure_set(x_616, 1, x_2); +x_629 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2; +x_630 = l_Lean_Expr_isConstOf(x_3, x_629); +if (x_630 == 0) +{ +lean_object* x_631; +x_631 = lean_box(0); +x_617 = x_631; +goto block_628; +} +else +{ +lean_object* x_632; lean_object* x_633; uint8_t x_634; +x_632 = lean_array_get_size(x_4); +x_633 = lean_unsigned_to_nat(2u); +x_634 = lean_nat_dec_eq(x_632, x_633); +if (x_634 == 0) +{ +lean_object* x_635; +lean_dec(x_632); +x_635 = lean_box(0); +x_617 = x_635; +goto block_628; +} +else +{ +lean_object* x_636; uint8_t x_637; +lean_dec(x_616); +lean_dec(x_3); +x_636 = lean_unsigned_to_nat(0u); +x_637 = lean_nat_dec_lt(x_636, x_632); +lean_dec(x_632); +if (x_637 == 0) +{ +lean_object* x_638; lean_object* x_639; lean_object* x_640; +lean_dec(x_4); +x_638 = l_Lean_instInhabitedExpr; +x_639 = l_outOfBounds___rarg(x_638); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_639); +x_640 = l_Lean_Meta_Grind_internalize(x_639, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_615); +if (lean_obj_tag(x_640) == 0) +{ +lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; +x_641 = lean_ctor_get(x_640, 1); +lean_inc(x_641); +lean_dec(x_640); +lean_inc(x_1); +x_642 = l_Lean_Meta_Grind_registerParent(x_1, x_639, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_641); +lean_dec(x_639); +x_643 = lean_ctor_get(x_642, 0); +lean_inc(x_643); +x_644 = lean_ctor_get(x_642, 1); +lean_inc(x_644); +lean_dec(x_642); +x_645 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_643, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_644); +lean_dec(x_643); +return x_645; +} +else +{ +uint8_t x_646; +lean_dec(x_639); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_646 = !lean_is_exclusive(x_640); +if (x_646 == 0) +{ +return x_640; +} +else +{ +lean_object* x_647; lean_object* x_648; lean_object* x_649; +x_647 = lean_ctor_get(x_640, 0); +x_648 = lean_ctor_get(x_640, 1); +lean_inc(x_648); +lean_inc(x_647); +lean_dec(x_640); +x_649 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_649, 0, x_647); +lean_ctor_set(x_649, 1, x_648); +return x_649; +} +} +} +else +{ +lean_object* x_650; lean_object* x_651; +x_650 = lean_array_fget(x_4, x_636); +lean_dec(x_4); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_650); +x_651 = l_Lean_Meta_Grind_internalize(x_650, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_615); +if (lean_obj_tag(x_651) == 0) +{ +lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; +x_652 = lean_ctor_get(x_651, 1); +lean_inc(x_652); +lean_dec(x_651); +lean_inc(x_1); +x_653 = l_Lean_Meta_Grind_registerParent(x_1, x_650, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_652); +lean_dec(x_650); +x_654 = lean_ctor_get(x_653, 0); +lean_inc(x_654); +x_655 = lean_ctor_get(x_653, 1); +lean_inc(x_655); +lean_dec(x_653); +x_656 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__1(x_1, x_2, x_654, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_655); +lean_dec(x_654); +return x_656; +} +else +{ +uint8_t x_657; +lean_dec(x_650); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_657 = !lean_is_exclusive(x_651); +if (x_657 == 0) +{ +return x_651; +} +else +{ +lean_object* x_658; lean_object* x_659; lean_object* x_660; +x_658 = lean_ctor_get(x_651, 0); +x_659 = lean_ctor_get(x_651, 1); +lean_inc(x_659); +lean_inc(x_658); +lean_dec(x_651); +x_660 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_660, 0, x_658); +lean_ctor_set(x_660, 1, x_659); +return x_660; +} +} +} +} +} +block_628: +{ +lean_object* x_618; +lean_dec(x_617); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_3); +x_618 = l_Lean_Meta_Grind_internalize(x_3, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_615); +if (lean_obj_tag(x_618) == 0) +{ +lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; +x_619 = lean_ctor_get(x_618, 1); +lean_inc(x_619); +lean_dec(x_618); +lean_inc(x_1); +x_620 = l_Lean_Meta_Grind_registerParent(x_1, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_619); +lean_dec(x_3); +x_621 = lean_ctor_get(x_620, 0); +lean_inc(x_621); +x_622 = lean_ctor_get(x_620, 1); +lean_inc(x_622); +lean_dec(x_620); +x_623 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___lambda__2(x_4, x_1, x_2, x_616, x_621, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_622); +lean_dec(x_621); +lean_dec(x_4); +return x_623; +} +else +{ +uint8_t x_624; +lean_dec(x_616); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_624 = !lean_is_exclusive(x_618); +if (x_624 == 0) +{ +return x_618; +} +else +{ +lean_object* x_625; lean_object* x_626; lean_object* x_627; +x_625 = lean_ctor_get(x_618, 0); +x_626 = lean_ctor_get(x_618, 1); +lean_inc(x_626); +lean_inc(x_625); +lean_dec(x_618); +x_627 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_627, 0, x_625); +lean_ctor_set(x_627, 1, x_626); +return x_627; +} +} +} +} +else +{ +uint8_t x_661; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_661 = !lean_is_exclusive(x_614); +if (x_661 == 0) +{ +return x_614; +} +else +{ +lean_object* x_662; lean_object* x_663; lean_object* x_664; +x_662 = lean_ctor_get(x_614, 0); +x_663 = lean_ctor_get(x_614, 1); +lean_inc(x_663); +lean_inc(x_662); +lean_dec(x_614); +x_664 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_664, 0, x_662); +lean_ctor_set(x_664, 1, x_663); +return x_664; +} +} +} +else +{ +uint8_t x_665; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_665 = !lean_is_exclusive(x_610); +if (x_665 == 0) +{ +return x_610; +} +else +{ +lean_object* x_666; lean_object* x_667; lean_object* x_668; +x_666 = lean_ctor_get(x_610, 0); +x_667 = lean_ctor_get(x_610, 1); +lean_inc(x_667); +lean_inc(x_666); +lean_dec(x_610); +x_668 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_668, 0, x_666); +lean_ctor_set(x_668, 1, x_667); +return x_668; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = 0; +x_14 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_13, x_13, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_propagateUp(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Internalize", 34, 34); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.internalize", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_internalize___lambda__3___closed__1; +x_2 = l_Lean_Meta_Grind_internalize___lambda__3___closed__2; +x_3 = lean_unsigned_to_nat(141u); +x_4 = lean_unsigned_to_nat(16u); +x_5 = l_Lean_Meta_Grind_internalize___lambda__3___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unexpected term during internalization", 38, 38); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_internalize___lambda__3___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_2); +lean_dec(x_1); +x_13 = l_Lean_Meta_Grind_internalize___lambda__3___closed__4; +x_14 = l_panic___at_Lean_Meta_Grind_internalize___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +case 2: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = 0; +x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_20, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_21; +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 1); +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +x_25 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +lean_inc(x_1); +x_27 = l_Lean_indentExpr(x_1); +x_28 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +lean_ctor_set_tag(x_16, 7); +lean_ctor_set(x_16, 1, x_27); +lean_ctor_set(x_16, 0, x_28); +x_29 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_16); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_30, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = 0; +x_34 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_33, x_33, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_34; +} +else +{ +uint8_t x_35; +lean_free_object(x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_25); +if (x_35 == 0) +{ +return x_25; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_25, 0); +x_37 = lean_ctor_get(x_25, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_25); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +else +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_16, 1); +lean_inc(x_39); +lean_dec(x_16); +x_40 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +lean_inc(x_1); +x_42 = l_Lean_indentExpr(x_1); +x_43 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +x_45 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_46, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_41); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = 0; +x_50 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_49, x_49, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_48); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_50; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_51 = lean_ctor_get(x_40, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_40, 1); +lean_inc(x_52); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + x_53 = x_40; +} else { + lean_dec_ref(x_40); + x_53 = lean_box(0); +} +if (lean_is_scalar(x_53)) { + x_54 = lean_alloc_ctor(1, 2, 0); +} else { + x_54 = x_53; +} +lean_ctor_set(x_54, 0, x_51); +lean_ctor_set(x_54, 1, x_52); +return x_54; +} +} +} +} +case 3: +{ +lean_object* x_55; lean_object* x_56; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_55 = lean_box(0); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_12); +return x_56; +} +case 4: +{ +lean_object* x_57; +x_57 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_57; +} +case 5: +{ +lean_object* x_58; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_58 = l_Lean_Meta_isLitValue(x_1, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; uint8_t x_60; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_unbox(x_59); +lean_dec(x_59); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_unsigned_to_nat(0u); +x_63 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_62); +x_64 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1; +lean_inc(x_63); +x_65 = lean_mk_array(x_63, x_64); +x_66 = lean_unsigned_to_nat(1u); +x_67 = lean_nat_sub(x_63, x_66); +lean_dec(x_63); +lean_inc(x_1); +x_68 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3(x_1, x_2, x_1, x_65, x_67, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_61); +return x_68; +} +else +{ +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_58, 1); +lean_inc(x_69); +lean_dec(x_58); +x_70 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_69); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_70; +} +} +else +{ +uint8_t x_71; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_71 = !lean_is_exclusive(x_58); +if (x_71 == 0) +{ +return x_58; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_58, 0); +x_73 = lean_ctor_get(x_58, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_58); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +case 7: +{ +lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; lean_object* x_105; +x_75 = lean_ctor_get(x_1, 1); +lean_inc(x_75); +x_76 = lean_ctor_get(x_1, 2); +lean_inc(x_76); +x_77 = 0; +lean_inc(x_2); +lean_inc(x_1); +x_78 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_77, x_77, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_80 = x_78; +} else { + lean_dec_ref(x_78); + x_80 = lean_box(0); +} +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_75); +x_105 = l_Lean_Meta_isProp(x_75, x_8, x_9, x_10, x_11, x_79); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; uint8_t x_107; +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_unbox(x_106); +if (x_107 == 0) +{ +lean_object* x_108; uint8_t x_109; +x_108 = lean_ctor_get(x_105, 1); +lean_inc(x_108); +lean_dec(x_105); +x_109 = lean_unbox(x_106); +lean_dec(x_106); +x_81 = x_109; +x_82 = x_108; +goto block_104; +} +else +{ +lean_object* x_110; lean_object* x_111; +lean_dec(x_106); +x_110 = lean_ctor_get(x_105, 1); +lean_inc(x_110); +lean_dec(x_105); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_111 = l_Lean_Meta_isProp(x_1, x_8, x_9, x_10, x_11, x_110); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; uint8_t x_114; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = lean_unbox(x_112); +lean_dec(x_112); +x_81 = x_114; +x_82 = x_113; +goto block_104; +} +else +{ +uint8_t x_115; +lean_dec(x_80); +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_115 = !lean_is_exclusive(x_111); +if (x_115 == 0) +{ +return x_111; +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_111, 0); +x_117 = lean_ctor_get(x_111, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_111); +x_118 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +return x_118; +} +} +} +} +else +{ +uint8_t x_119; +lean_dec(x_80); +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_119 = !lean_is_exclusive(x_105); +if (x_119 == 0) +{ +return x_105; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_105, 0); +x_121 = lean_ctor_get(x_105, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_105); +x_122 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +return x_122; +} +} +block_104: +{ +if (x_81 == 0) +{ +lean_object* x_83; lean_object* x_84; +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_83 = lean_box(0); +if (lean_is_scalar(x_80)) { + x_84 = lean_alloc_ctor(0, 2, 0); +} else { + x_84 = x_80; +} +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_82); +return x_84; +} +else +{ +lean_object* x_85; +lean_dec(x_80); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_75); +x_85 = l_Lean_Meta_Grind_internalize(x_75, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_82); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +lean_dec(x_85); +lean_inc(x_1); +x_87 = l_Lean_Meta_Grind_registerParent(x_1, x_75, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_86); +lean_dec(x_75); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = l_Lean_Expr_hasLooseBVars(x_76); +if (x_89 == 0) +{ +lean_object* x_90; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_76); +x_90 = l_Lean_Meta_Grind_internalize(x_76, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_88); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_91 = lean_ctor_get(x_90, 1); +lean_inc(x_91); +lean_dec(x_90); +lean_inc(x_1); +x_92 = l_Lean_Meta_Grind_registerParent(x_1, x_76, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_91); +lean_dec(x_76); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +lean_dec(x_92); +x_94 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_93); +return x_94; +} +else +{ +uint8_t x_95; +lean_dec(x_76); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_95 = !lean_is_exclusive(x_90); +if (x_95 == 0) +{ +return x_90; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_90, 0); +x_97 = lean_ctor_get(x_90, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_90); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +else +{ +lean_object* x_99; +lean_dec(x_76); +lean_dec(x_2); +x_99 = l_Lean_Meta_Grind_propagateUp(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_88); +return x_99; +} +} +else +{ +uint8_t x_100; +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_100 = !lean_is_exclusive(x_85); +if (x_100 == 0) +{ +return x_85; +} +else +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_85, 0); +x_102 = lean_ctor_get(x_85, 1); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_85); +x_103 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +return x_103; +} +} +} +} +} +case 9: +{ +lean_object* x_123; +x_123 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_123; +} +case 10: +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; +x_124 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_125 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_124, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_unbox(x_126); +lean_dec(x_126); +if (x_127 == 0) +{ +lean_object* x_128; uint8_t x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_125, 1); +lean_inc(x_128); +lean_dec(x_125); +x_129 = 0; +x_130 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_129, x_129, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_128); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_130; +} +else +{ +uint8_t x_131; +x_131 = !lean_is_exclusive(x_125); +if (x_131 == 0) +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_125, 1); +x_133 = lean_ctor_get(x_125, 0); +lean_dec(x_133); +x_134 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_132); +if (lean_obj_tag(x_134) == 0) +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +lean_dec(x_134); +lean_inc(x_1); +x_136 = l_Lean_indentExpr(x_1); +x_137 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +lean_ctor_set_tag(x_125, 7); +lean_ctor_set(x_125, 1, x_136); +lean_ctor_set(x_125, 0, x_137); +x_138 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_139 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_139, 0, x_125); +lean_ctor_set(x_139, 1, x_138); +x_140 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_124, x_139, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_135); +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +lean_dec(x_140); +x_142 = 0; +x_143 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_142, x_142, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_141); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_143; +} +else +{ +uint8_t x_144; +lean_free_object(x_125); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_144 = !lean_is_exclusive(x_134); +if (x_144 == 0) +{ +return x_134; +} +else +{ +lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_134, 0); +x_146 = lean_ctor_get(x_134, 1); +lean_inc(x_146); +lean_inc(x_145); +lean_dec(x_134); +x_147 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +return x_147; +} +} +} +else +{ +lean_object* x_148; lean_object* x_149; +x_148 = lean_ctor_get(x_125, 1); +lean_inc(x_148); +lean_dec(x_125); +x_149 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_148); +if (lean_obj_tag(x_149) == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; lean_object* x_159; +x_150 = lean_ctor_get(x_149, 1); +lean_inc(x_150); +lean_dec(x_149); +lean_inc(x_1); +x_151 = l_Lean_indentExpr(x_1); +x_152 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +x_153 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_151); +x_154 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_155 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_155, 0, x_153); +lean_ctor_set(x_155, 1, x_154); +x_156 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_124, x_155, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_150); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +lean_dec(x_156); +x_158 = 0; +x_159 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_158, x_158, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_157); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_159; +} +else +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_160 = lean_ctor_get(x_149, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_149, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_149)) { + lean_ctor_release(x_149, 0); + lean_ctor_release(x_149, 1); + x_162 = x_149; +} else { + lean_dec_ref(x_149); + x_162 = lean_box(0); +} +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(1, 2, 0); +} else { + x_163 = x_162; +} +lean_ctor_set(x_163, 0, x_160); +lean_ctor_set(x_163, 1, x_161); +return x_163; +} +} +} +} +case 11: +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; +x_164 = l_Lean_Meta_Grind_addCongrTable___closed__2; +x_165 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_164, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_166 = lean_ctor_get(x_165, 0); +lean_inc(x_166); +x_167 = lean_unbox(x_166); +lean_dec(x_166); +if (x_167 == 0) +{ +lean_object* x_168; uint8_t x_169; lean_object* x_170; +x_168 = lean_ctor_get(x_165, 1); +lean_inc(x_168); +lean_dec(x_165); +x_169 = 0; +x_170 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_169, x_169, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_168); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_170; +} +else +{ +uint8_t x_171; +x_171 = !lean_is_exclusive(x_165); +if (x_171 == 0) +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_172 = lean_ctor_get(x_165, 1); +x_173 = lean_ctor_get(x_165, 0); +lean_dec(x_173); +x_174 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_172); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; lean_object* x_183; +x_175 = lean_ctor_get(x_174, 1); +lean_inc(x_175); +lean_dec(x_174); +lean_inc(x_1); +x_176 = l_Lean_indentExpr(x_1); +x_177 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +lean_ctor_set_tag(x_165, 7); +lean_ctor_set(x_165, 1, x_176); +lean_ctor_set(x_165, 0, x_177); +x_178 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_179 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_179, 0, x_165); +lean_ctor_set(x_179, 1, x_178); +x_180 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_164, x_179, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_175); +x_181 = lean_ctor_get(x_180, 1); +lean_inc(x_181); +lean_dec(x_180); +x_182 = 0; +x_183 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_182, x_182, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_181); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_183; +} +else +{ +uint8_t x_184; +lean_free_object(x_165); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_184 = !lean_is_exclusive(x_174); +if (x_184 == 0) +{ +return x_174; +} +else +{ +lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_185 = lean_ctor_get(x_174, 0); +x_186 = lean_ctor_get(x_174, 1); +lean_inc(x_186); +lean_inc(x_185); +lean_dec(x_174); +x_187 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_187, 0, x_185); +lean_ctor_set(x_187, 1, x_186); +return x_187; +} +} +} +else +{ +lean_object* x_188; lean_object* x_189; +x_188 = lean_ctor_get(x_165, 1); +lean_inc(x_188); +lean_dec(x_165); +x_189 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_188); +if (lean_obj_tag(x_189) == 0) +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; uint8_t x_198; lean_object* x_199; +x_190 = lean_ctor_get(x_189, 1); +lean_inc(x_190); +lean_dec(x_189); +lean_inc(x_1); +x_191 = l_Lean_indentExpr(x_1); +x_192 = l_Lean_Meta_Grind_internalize___lambda__3___closed__6; +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_191); +x_194 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_164, x_195, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_190); +x_197 = lean_ctor_get(x_196, 1); +lean_inc(x_197); +lean_dec(x_196); +x_198 = 0; +x_199 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_198, x_198, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_197); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_199; +} +else +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_200 = lean_ctor_get(x_189, 0); +lean_inc(x_200); +x_201 = lean_ctor_get(x_189, 1); +lean_inc(x_201); +if (lean_is_exclusive(x_189)) { + lean_ctor_release(x_189, 0); + lean_ctor_release(x_189, 1); + x_202 = x_189; +} else { + lean_dec_ref(x_189); + x_202 = lean_box(0); +} +if (lean_is_scalar(x_202)) { + x_203 = lean_alloc_ctor(1, 2, 0); +} else { + x_203 = x_202; +} +lean_ctor_set(x_203, 0, x_200); +lean_ctor_set(x_203, 1, x_201); +return x_203; +} +} +} +} +default: +{ +uint8_t x_204; lean_object* x_205; +x_204 = 0; +x_205 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_204, x_204, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_205; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("internalize", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_internalize___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; +x_2 = l_Lean_Meta_Grind_internalize___lambda__4___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = l_Lean_Meta_Grind_internalize___lambda__4___closed__2; +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_box(0); +x_19 = l_Lean_Meta_Grind_internalize___lambda__3(x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +return x_19; +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_14); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 1); +x_22 = lean_ctor_get(x_14, 0); +lean_dec(x_22); +x_23 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +lean_inc(x_1); +x_25 = l_Lean_MessageData_ofExpr(x_1); +x_26 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +lean_ctor_set_tag(x_14, 7); +lean_ctor_set(x_14, 1, x_25); +lean_ctor_set(x_14, 0, x_26); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Meta_Grind_internalize___lambda__3(x_1, x_2, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_29); +return x_31; +} +else +{ +uint8_t x_32; +lean_free_object(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) +{ +return x_23; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_23, 0); +x_34 = lean_ctor_get(x_23, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_23); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); +lean_dec(x_14); +x_37 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_36); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +lean_inc(x_1); +x_39 = l_Lean_MessageData_ofExpr(x_1); +x_40 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_internalize___lambda__3(x_1, x_2, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); +lean_dec(x_44); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_47 = lean_ctor_get(x_37, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_37, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_49 = x_37; +} else { + lean_dec_ref(x_37); + x_49 = lean_box(0); +} +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(1, 2, 0); +} else { + x_50 = x_49; +} +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_alreadyInternalized(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l_Lean_Meta_Grind_internalize___lambda__4(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_12, 0, x_20); +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; +} +else +{ +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_69_(x_3, x_11); +lean_dec(x_11); +return x_12; +} +case 1: +{ +lean_object* x_13; size_t x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_usize_shift_right(x_2, x_5); +x_1 = x_13; +x_2 = x_14; +goto _start; +} +default: +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 1); +lean_inc(x_18); +lean_dec(x_1); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_17, x_18, lean_box(0), x_19, x_3); +lean_dec(x_18); +lean_dec(x_17); +return x_20; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_HeadIndex_hash(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = l_List_reverse___rarg(x_3); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +x_8 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_6); +if (x_8 == 0) +{ +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_free_object(x_2); +lean_dec(x_6); +x_2 = x_7; +goto _start; +} +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_2); +lean_inc(x_1); +x_13 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_11); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_3); +x_2 = x_12; +x_3 = x_14; +goto _start; +} +else +{ +lean_dec(x_11); +x_2 = x_12; +goto _start; +} +} +} +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_13, 10); +x_17 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_16, x_1); +lean_ctor_set(x_13, 10, x_17); +x_18 = lean_st_ref_set(x_3, x_13, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +x_28 = lean_ctor_get(x_13, 3); +x_29 = lean_ctor_get(x_13, 4); +x_30 = lean_ctor_get(x_13, 5); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*19); +x_32 = lean_ctor_get(x_13, 6); +x_33 = lean_ctor_get(x_13, 7); +x_34 = lean_ctor_get(x_13, 8); +x_35 = lean_ctor_get(x_13, 9); +x_36 = lean_ctor_get(x_13, 10); +x_37 = lean_ctor_get(x_13, 11); +x_38 = lean_ctor_get(x_13, 12); +x_39 = lean_ctor_get(x_13, 13); +x_40 = lean_ctor_get(x_13, 14); +x_41 = lean_ctor_get(x_13, 15); +x_42 = lean_ctor_get(x_13, 16); +x_43 = lean_ctor_get(x_13, 17); +x_44 = lean_ctor_get(x_13, 18); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_45 = l_Lean_Meta_Grind_EMatchTheorems_insert(x_36, x_1); +x_46 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_46, 0, x_25); +lean_ctor_set(x_46, 1, x_26); +lean_ctor_set(x_46, 2, x_27); +lean_ctor_set(x_46, 3, x_28); +lean_ctor_set(x_46, 4, x_29); +lean_ctor_set(x_46, 5, x_30); +lean_ctor_set(x_46, 6, x_32); +lean_ctor_set(x_46, 7, x_33); +lean_ctor_set(x_46, 8, x_34); +lean_ctor_set(x_46, 9, x_35); +lean_ctor_set(x_46, 10, x_45); +lean_ctor_set(x_46, 11, x_37); +lean_ctor_set(x_46, 12, x_38); +lean_ctor_set(x_46, 13, x_39); +lean_ctor_set(x_46, 14, x_40); +lean_ctor_set(x_46, 15, x_41); +lean_ctor_set(x_46, 16, x_42); +lean_ctor_set(x_46, 17, x_43); +lean_ctor_set(x_46, 18, x_44); +lean_ctor_set_uint8(x_46, sizeof(void*)*19, x_31); +x_47 = lean_st_ref_set(x_3, x_46, x_14); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); +} +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_49; +} +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ematch", 6, 6); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1; +x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("reinsert `", 10, 10); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_18; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_dec(x_7); +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_6, 1); +lean_inc(x_20); +lean_dec(x_6); +x_26 = lean_st_ref_get(x_9, x_17); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_ctor_get(x_27, 10); +lean_inc(x_29); +lean_dec(x_27); +x_30 = !lean_is_exclusive(x_19); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_31 = lean_ctor_get(x_19, 0); +x_32 = lean_ctor_get(x_19, 1); +x_33 = lean_ctor_get(x_19, 2); +x_34 = lean_ctor_get(x_19, 3); +x_35 = lean_ctor_get(x_19, 4); +x_36 = lean_ctor_get(x_19, 5); +x_37 = l_Lean_Meta_Grind_EMatchTheorems_isErased(x_29, x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_box(0); +lean_inc(x_3); +x_39 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_3, x_35, x_38); +lean_inc(x_36); +lean_inc(x_39); +lean_ctor_set(x_19, 4, x_39); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; +lean_dec(x_36); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_40 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem(x_19, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_21 = x_42; +x_22 = x_41; +goto block_25; +} +else +{ +uint8_t x_43; +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_40); +if (x_43 == 0) +{ +return x_40; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_40, 0); +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_40); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +uint8_t x_47; +x_47 = !lean_is_exclusive(x_39); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_48 = lean_ctor_get(x_39, 1); +lean_dec(x_48); +x_49 = lean_ctor_get(x_39, 0); +lean_dec(x_49); +x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_51 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_50, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_unbox(x_52); +lean_dec(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_free_object(x_39); +lean_dec(x_36); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); +x_55 = lean_box(0); +x_56 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_55, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_54); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_21 = x_57; +x_22 = x_58; +goto block_25; +} +else +{ +uint8_t x_59; +x_59 = !lean_is_exclusive(x_51); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_51, 1); +x_61 = lean_ctor_get(x_51, 0); +lean_dec(x_61); +x_62 = l_Lean_Meta_Grind_updateLastTag(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_60); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = l_Lean_Meta_Grind_Origin_key(x_36); +lean_dec(x_36); +x_65 = l_Lean_MessageData_ofName(x_64); +x_66 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; +lean_ctor_set_tag(x_51, 7); +lean_ctor_set(x_51, 1, x_65); +lean_ctor_set(x_51, 0, x_66); +x_67 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_67); +lean_ctor_set(x_39, 0, x_51); +x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_50, x_39, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_63); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_69, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_70); +lean_dec(x_69); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +x_21 = x_72; +x_22 = x_73; +goto block_25; +} +else +{ +uint8_t x_74; +lean_free_object(x_51); +lean_free_object(x_39); +lean_dec(x_19); +lean_dec(x_36); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_74 = !lean_is_exclusive(x_62); +if (x_74 == 0) +{ +return x_62; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_62, 0); +x_76 = lean_ctor_get(x_62, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_62); +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +} +} +else +{ +lean_object* x_78; lean_object* x_79; +x_78 = lean_ctor_get(x_51, 1); +lean_inc(x_78); +lean_dec(x_51); +x_79 = l_Lean_Meta_Grind_updateLastTag(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_78); +if (lean_obj_tag(x_79) == 0) +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = l_Lean_Meta_Grind_Origin_key(x_36); +lean_dec(x_36); +x_82 = l_Lean_MessageData_ofName(x_81); +x_83 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_82); +x_85 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_85); +lean_ctor_set(x_39, 0, x_84); +x_86 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_50, x_39, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_80); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_87, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_88); +lean_dec(x_87); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); +x_21 = x_90; +x_22 = x_91; +goto block_25; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_free_object(x_39); +lean_dec(x_19); +lean_dec(x_36); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_92 = lean_ctor_get(x_79, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_79, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_94 = x_79; +} else { + lean_dec_ref(x_79); + x_94 = lean_box(0); +} +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(1, 2, 0); +} else { + x_95 = x_94; +} +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +return x_95; +} +} +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +lean_dec(x_39); +x_96 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_97 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_96, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_unbox(x_98); +lean_dec(x_98); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_36); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +x_101 = lean_box(0); +x_102 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_101, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_100); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_21 = x_103; +x_22 = x_104; +goto block_25; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_97, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_106 = x_97; +} else { + lean_dec_ref(x_97); + x_106 = lean_box(0); +} +x_107 = l_Lean_Meta_Grind_updateLastTag(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_105); +if (lean_obj_tag(x_107) == 0) +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_108 = lean_ctor_get(x_107, 1); +lean_inc(x_108); +lean_dec(x_107); +x_109 = l_Lean_Meta_Grind_Origin_key(x_36); +lean_dec(x_36); +x_110 = l_Lean_MessageData_ofName(x_109); +x_111 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; +if (lean_is_scalar(x_106)) { + x_112 = lean_alloc_ctor(7, 2, 0); +} else { + x_112 = x_106; + lean_ctor_set_tag(x_112, 7); +} +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_110); +x_113 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +x_114 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +x_115 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_96, x_114, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_108); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_19, x_116, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_117); +lean_dec(x_116); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_21 = x_119; +x_22 = x_120; +goto block_25; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_106); +lean_dec(x_19); +lean_dec(x_36); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_121 = lean_ctor_get(x_107, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_107, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_123 = x_107; +} else { + lean_dec_ref(x_107); + x_123 = lean_box(0); +} +if (lean_is_scalar(x_123)) { + x_124 = lean_alloc_ctor(1, 2, 0); +} else { + x_124 = x_123; +} +lean_ctor_set(x_124, 0, x_121); +lean_ctor_set(x_124, 1, x_122); +return x_124; +} +} +} +} +} +else +{ +lean_object* x_125; +lean_free_object(x_19); +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +x_125 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_21 = x_125; +x_22 = x_28; +goto block_25; +} +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_126 = lean_ctor_get(x_19, 0); +x_127 = lean_ctor_get(x_19, 1); +x_128 = lean_ctor_get(x_19, 2); +x_129 = lean_ctor_get(x_19, 3); +x_130 = lean_ctor_get(x_19, 4); +x_131 = lean_ctor_get(x_19, 5); +lean_inc(x_131); +lean_inc(x_130); +lean_inc(x_129); +lean_inc(x_128); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_19); +x_132 = l_Lean_Meta_Grind_EMatchTheorems_isErased(x_29, x_131); +if (x_132 == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_box(0); +lean_inc(x_3); +x_134 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_3, x_130, x_133); +lean_inc(x_131); +lean_inc(x_134); +x_135 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_135, 0, x_126); +lean_ctor_set(x_135, 1, x_127); +lean_ctor_set(x_135, 2, x_128); +lean_ctor_set(x_135, 3, x_129); +lean_ctor_set(x_135, 4, x_134); +lean_ctor_set(x_135, 5, x_131); +if (lean_obj_tag(x_134) == 0) +{ +lean_object* x_136; +lean_dec(x_131); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_1); +x_136 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem(x_135, x_1, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_137; lean_object* x_138; +x_137 = lean_ctor_get(x_136, 1); +lean_inc(x_137); +lean_dec(x_136); +x_138 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_21 = x_138; +x_22 = x_137; +goto block_25; +} +else +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_139 = lean_ctor_get(x_136, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_136, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_141 = x_136; +} else { + lean_dec_ref(x_136); + x_141 = lean_box(0); +} +if (lean_is_scalar(x_141)) { + x_142 = lean_alloc_ctor(1, 2, 0); +} else { + x_142 = x_141; +} +lean_ctor_set(x_142, 0, x_139); +lean_ctor_set(x_142, 1, x_140); +return x_142; +} +} +else +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_143 = x_134; +} else { + lean_dec_ref(x_134); + x_143 = lean_box(0); +} +x_144 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_145 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_144, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_28); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_unbox(x_146); +lean_dec(x_146); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +lean_dec(x_143); +lean_dec(x_131); +x_148 = lean_ctor_get(x_145, 1); +lean_inc(x_148); +lean_dec(x_145); +x_149 = lean_box(0); +x_150 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_135, x_149, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_148); +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_21 = x_151; +x_22 = x_152; +goto block_25; +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_153 = lean_ctor_get(x_145, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_154 = x_145; +} else { + lean_dec_ref(x_145); + x_154 = lean_box(0); +} +x_155 = l_Lean_Meta_Grind_updateLastTag(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_153); +if (lean_obj_tag(x_155) == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_156 = lean_ctor_get(x_155, 1); +lean_inc(x_156); +lean_dec(x_155); +x_157 = l_Lean_Meta_Grind_Origin_key(x_131); +lean_dec(x_131); +x_158 = l_Lean_MessageData_ofName(x_157); +x_159 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4; +if (lean_is_scalar(x_154)) { + x_160 = lean_alloc_ctor(7, 2, 0); +} else { + x_160 = x_154; + lean_ctor_set_tag(x_160, 7); +} +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_158); +x_161 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6; +if (lean_is_scalar(x_143)) { + x_162 = lean_alloc_ctor(7, 2, 0); +} else { + x_162 = x_143; + lean_ctor_set_tag(x_162, 7); +} +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +x_163 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_144, x_162, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_156); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec(x_163); +x_166 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_135, x_164, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_165); +lean_dec(x_164); +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec(x_166); +x_21 = x_167; +x_22 = x_168; +goto block_25; +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +lean_dec(x_154); +lean_dec(x_143); +lean_dec(x_135); +lean_dec(x_131); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +lean_dec(x_1); +x_169 = lean_ctor_get(x_155, 0); +lean_inc(x_169); +x_170 = lean_ctor_get(x_155, 1); +lean_inc(x_170); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_171 = x_155; +} else { + lean_dec_ref(x_155); + x_171 = lean_box(0); +} +if (lean_is_scalar(x_171)) { + x_172 = lean_alloc_ctor(1, 2, 0); +} else { + x_172 = x_171; +} +lean_ctor_set(x_172, 0, x_169); +lean_ctor_set(x_172, 1, x_170); +return x_172; +} +} +} +} +else +{ +lean_object* x_173; +lean_dec(x_131); +lean_dec(x_130); +lean_dec(x_129); +lean_dec(x_128); +lean_dec(x_127); +lean_dec(x_126); +x_173 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1; +x_21 = x_173; +x_22 = x_28; +goto block_25; +} +} +block_25: +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +x_6 = x_20; +x_7 = x_23; +x_8 = lean_box(0); +x_17 = x_22; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_13, 10); +lean_inc(x_14); +lean_dec(x_13); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_12, 1); +x_17 = lean_ctor_get(x_12, 0); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_14, 0); +x_20 = lean_ctor_get(x_14, 1); +x_21 = lean_ctor_get(x_14, 2); +lean_inc(x_19); +x_22 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_19, x_1); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; +lean_free_object(x_14); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_box(0); +lean_ctor_set(x_12, 0, x_23); +return x_12; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +lean_free_object(x_12); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_19, x_1); +lean_ctor_set(x_14, 0, x_25); +x_26 = lean_st_ref_take(x_3, x_16); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = !lean_is_exclusive(x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_30 = lean_ctor_get(x_27, 10); +lean_dec(x_30); +lean_ctor_set(x_27, 10, x_14); +x_31 = lean_st_ref_set(x_3, x_27, x_28); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_st_ref_get(x_3, x_32); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_ctor_get(x_34, 4); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_box(0); +x_38 = lean_box(0); +lean_inc(x_24); +x_39 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_2, x_24, x_36, x_37, x_24, x_24, x_38, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +lean_dec(x_24); +if (lean_obj_tag(x_39) == 0) +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_39, 0); +lean_dec(x_41); +lean_ctor_set(x_39, 0, x_38); +return x_39; +} +else +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); +lean_dec(x_39); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +else +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_39); +if (x_44 == 0) +{ +return x_39; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 0); +x_46 = lean_ctor_get(x_39, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_39); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_48 = lean_ctor_get(x_27, 0); +x_49 = lean_ctor_get(x_27, 1); +x_50 = lean_ctor_get(x_27, 2); +x_51 = lean_ctor_get(x_27, 3); +x_52 = lean_ctor_get(x_27, 4); +x_53 = lean_ctor_get(x_27, 5); +x_54 = lean_ctor_get_uint8(x_27, sizeof(void*)*19); +x_55 = lean_ctor_get(x_27, 6); +x_56 = lean_ctor_get(x_27, 7); +x_57 = lean_ctor_get(x_27, 8); +x_58 = lean_ctor_get(x_27, 9); +x_59 = lean_ctor_get(x_27, 11); +x_60 = lean_ctor_get(x_27, 12); +x_61 = lean_ctor_get(x_27, 13); +x_62 = lean_ctor_get(x_27, 14); +x_63 = lean_ctor_get(x_27, 15); +x_64 = lean_ctor_get(x_27, 16); +x_65 = lean_ctor_get(x_27, 17); +x_66 = lean_ctor_get(x_27, 18); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_27); +x_67 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_67, 0, x_48); +lean_ctor_set(x_67, 1, x_49); +lean_ctor_set(x_67, 2, x_50); +lean_ctor_set(x_67, 3, x_51); +lean_ctor_set(x_67, 4, x_52); +lean_ctor_set(x_67, 5, x_53); +lean_ctor_set(x_67, 6, x_55); +lean_ctor_set(x_67, 7, x_56); +lean_ctor_set(x_67, 8, x_57); +lean_ctor_set(x_67, 9, x_58); +lean_ctor_set(x_67, 10, x_14); +lean_ctor_set(x_67, 11, x_59); +lean_ctor_set(x_67, 12, x_60); +lean_ctor_set(x_67, 13, x_61); +lean_ctor_set(x_67, 14, x_62); +lean_ctor_set(x_67, 15, x_63); +lean_ctor_set(x_67, 16, x_64); +lean_ctor_set(x_67, 17, x_65); +lean_ctor_set(x_67, 18, x_66); +lean_ctor_set_uint8(x_67, sizeof(void*)*19, x_54); +x_68 = lean_st_ref_set(x_3, x_67, x_28); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = lean_st_ref_get(x_3, x_69); +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); +x_73 = lean_ctor_get(x_71, 4); +lean_inc(x_73); +lean_dec(x_71); +x_74 = lean_box(0); +x_75 = lean_box(0); +lean_inc(x_24); +x_76 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_2, x_24, x_73, x_74, x_24, x_24, x_75, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_72); +lean_dec(x_24); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_78 = x_76; +} else { + lean_dec_ref(x_76); + x_78 = lean_box(0); +} +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_78; +} +lean_ctor_set(x_79, 0, x_75); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_80 = lean_ctor_get(x_76, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_76, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_82 = x_76; +} else { + lean_dec_ref(x_76); + x_82 = lean_box(0); +} +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(1, 2, 0); +} else { + x_83 = x_82; +} +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_81); +return x_83; +} +} +} +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_ctor_get(x_14, 0); +x_85 = lean_ctor_get(x_14, 1); +x_86 = lean_ctor_get(x_14, 2); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_14); +lean_inc(x_84); +x_87 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_84, x_1); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; +lean_dec(x_86); +lean_dec(x_85); +lean_dec(x_84); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +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_88 = lean_box(0); +lean_ctor_set(x_12, 0, x_88); +return x_12; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_free_object(x_12); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_84, x_1); +x_91 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_85); +lean_ctor_set(x_91, 2, x_86); +x_92 = lean_st_ref_take(x_3, x_16); +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_93, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_93, 1); +lean_inc(x_96); +x_97 = lean_ctor_get(x_93, 2); +lean_inc(x_97); +x_98 = lean_ctor_get(x_93, 3); +lean_inc(x_98); +x_99 = lean_ctor_get(x_93, 4); +lean_inc(x_99); +x_100 = lean_ctor_get(x_93, 5); +lean_inc(x_100); +x_101 = lean_ctor_get_uint8(x_93, sizeof(void*)*19); +x_102 = lean_ctor_get(x_93, 6); +lean_inc(x_102); +x_103 = lean_ctor_get(x_93, 7); +lean_inc(x_103); +x_104 = lean_ctor_get(x_93, 8); +lean_inc(x_104); +x_105 = lean_ctor_get(x_93, 9); +lean_inc(x_105); +x_106 = lean_ctor_get(x_93, 11); +lean_inc(x_106); +x_107 = lean_ctor_get(x_93, 12); +lean_inc(x_107); +x_108 = lean_ctor_get(x_93, 13); +lean_inc(x_108); +x_109 = lean_ctor_get(x_93, 14); +lean_inc(x_109); +x_110 = lean_ctor_get(x_93, 15); +lean_inc(x_110); +x_111 = lean_ctor_get(x_93, 16); +lean_inc(x_111); +x_112 = lean_ctor_get(x_93, 17); +lean_inc(x_112); +x_113 = lean_ctor_get(x_93, 18); +lean_inc(x_113); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + lean_ctor_release(x_93, 2); + lean_ctor_release(x_93, 3); + lean_ctor_release(x_93, 4); + lean_ctor_release(x_93, 5); + lean_ctor_release(x_93, 6); + lean_ctor_release(x_93, 7); + lean_ctor_release(x_93, 8); + lean_ctor_release(x_93, 9); + lean_ctor_release(x_93, 10); + lean_ctor_release(x_93, 11); + lean_ctor_release(x_93, 12); + lean_ctor_release(x_93, 13); + lean_ctor_release(x_93, 14); + lean_ctor_release(x_93, 15); + lean_ctor_release(x_93, 16); + lean_ctor_release(x_93, 17); + lean_ctor_release(x_93, 18); + x_114 = x_93; +} else { + lean_dec_ref(x_93); + x_114 = lean_box(0); +} +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(0, 19, 1); +} else { + x_115 = x_114; +} +lean_ctor_set(x_115, 0, x_95); +lean_ctor_set(x_115, 1, x_96); +lean_ctor_set(x_115, 2, x_97); +lean_ctor_set(x_115, 3, x_98); +lean_ctor_set(x_115, 4, x_99); +lean_ctor_set(x_115, 5, x_100); +lean_ctor_set(x_115, 6, x_102); +lean_ctor_set(x_115, 7, x_103); +lean_ctor_set(x_115, 8, x_104); +lean_ctor_set(x_115, 9, x_105); +lean_ctor_set(x_115, 10, x_91); +lean_ctor_set(x_115, 11, x_106); +lean_ctor_set(x_115, 12, x_107); +lean_ctor_set(x_115, 13, x_108); +lean_ctor_set(x_115, 14, x_109); +lean_ctor_set(x_115, 15, x_110); +lean_ctor_set(x_115, 16, x_111); +lean_ctor_set(x_115, 17, x_112); +lean_ctor_set(x_115, 18, x_113); +lean_ctor_set_uint8(x_115, sizeof(void*)*19, x_101); +x_116 = lean_st_ref_set(x_3, x_115, x_94); +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +lean_dec(x_116); +x_118 = lean_st_ref_get(x_3, x_117); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = lean_ctor_get(x_119, 4); +lean_inc(x_121); +lean_dec(x_119); +x_122 = lean_box(0); +x_123 = lean_box(0); +lean_inc(x_89); +x_124 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_2, x_89, x_121, x_122, x_89, x_89, x_123, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_120); +lean_dec(x_89); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_126 = x_124; +} else { + lean_dec_ref(x_124); + x_126 = lean_box(0); +} +if (lean_is_scalar(x_126)) { + x_127 = lean_alloc_ctor(0, 2, 0); +} else { + x_127 = x_126; +} +lean_ctor_set(x_127, 0, x_123); +lean_ctor_set(x_127, 1, x_125); +return x_127; +} +else +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_128 = lean_ctor_get(x_124, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_124, 1); +lean_inc(x_129); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_130 = x_124; +} else { + lean_dec_ref(x_124); + x_130 = lean_box(0); +} +if (lean_is_scalar(x_130)) { + x_131 = lean_alloc_ctor(1, 2, 0); +} else { + x_131 = x_130; +} +lean_ctor_set(x_131, 0, x_128); +lean_ctor_set(x_131, 1, x_129); +return x_131; +} +} +} +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_132 = lean_ctor_get(x_12, 1); +lean_inc(x_132); +lean_dec(x_12); +x_133 = lean_ctor_get(x_14, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_14, 1); +lean_inc(x_134); +x_135 = lean_ctor_get(x_14, 2); +lean_inc(x_135); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + lean_ctor_release(x_14, 2); + x_136 = x_14; +} else { + lean_dec_ref(x_14); + x_136 = lean_box(0); +} +lean_inc(x_133); +x_137 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__9(x_133, x_1); +if (lean_obj_tag(x_137) == 0) +{ +lean_object* x_138; lean_object* x_139; +lean_dec(x_136); +lean_dec(x_135); +lean_dec(x_134); +lean_dec(x_133); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +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_138 = lean_box(0); +x_139 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_132); +return x_139; +} +else +{ +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; uint8_t x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_140 = lean_ctor_get(x_137, 0); +lean_inc(x_140); +lean_dec(x_137); +x_141 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_EMatchTheorems_retrieve_x3f___spec__1(x_133, x_1); +if (lean_is_scalar(x_136)) { + x_142 = lean_alloc_ctor(0, 3, 0); +} else { + x_142 = x_136; +} +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_134); +lean_ctor_set(x_142, 2, x_135); +x_143 = lean_st_ref_take(x_3, x_132); +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +x_146 = lean_ctor_get(x_144, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_144, 1); +lean_inc(x_147); +x_148 = lean_ctor_get(x_144, 2); +lean_inc(x_148); +x_149 = lean_ctor_get(x_144, 3); +lean_inc(x_149); +x_150 = lean_ctor_get(x_144, 4); +lean_inc(x_150); +x_151 = lean_ctor_get(x_144, 5); +lean_inc(x_151); +x_152 = lean_ctor_get_uint8(x_144, sizeof(void*)*19); +x_153 = lean_ctor_get(x_144, 6); +lean_inc(x_153); +x_154 = lean_ctor_get(x_144, 7); +lean_inc(x_154); +x_155 = lean_ctor_get(x_144, 8); +lean_inc(x_155); +x_156 = lean_ctor_get(x_144, 9); +lean_inc(x_156); +x_157 = lean_ctor_get(x_144, 11); +lean_inc(x_157); +x_158 = lean_ctor_get(x_144, 12); +lean_inc(x_158); +x_159 = lean_ctor_get(x_144, 13); +lean_inc(x_159); +x_160 = lean_ctor_get(x_144, 14); +lean_inc(x_160); +x_161 = lean_ctor_get(x_144, 15); +lean_inc(x_161); +x_162 = lean_ctor_get(x_144, 16); +lean_inc(x_162); +x_163 = lean_ctor_get(x_144, 17); +lean_inc(x_163); +x_164 = lean_ctor_get(x_144, 18); +lean_inc(x_164); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + lean_ctor_release(x_144, 2); + lean_ctor_release(x_144, 3); + lean_ctor_release(x_144, 4); + lean_ctor_release(x_144, 5); + lean_ctor_release(x_144, 6); + lean_ctor_release(x_144, 7); + lean_ctor_release(x_144, 8); + lean_ctor_release(x_144, 9); + lean_ctor_release(x_144, 10); + lean_ctor_release(x_144, 11); + lean_ctor_release(x_144, 12); + lean_ctor_release(x_144, 13); + lean_ctor_release(x_144, 14); + lean_ctor_release(x_144, 15); + lean_ctor_release(x_144, 16); + lean_ctor_release(x_144, 17); + lean_ctor_release(x_144, 18); + x_165 = x_144; +} else { + lean_dec_ref(x_144); + x_165 = lean_box(0); +} +if (lean_is_scalar(x_165)) { + x_166 = lean_alloc_ctor(0, 19, 1); +} else { + x_166 = x_165; +} +lean_ctor_set(x_166, 0, x_146); +lean_ctor_set(x_166, 1, x_147); +lean_ctor_set(x_166, 2, x_148); +lean_ctor_set(x_166, 3, x_149); +lean_ctor_set(x_166, 4, x_150); +lean_ctor_set(x_166, 5, x_151); +lean_ctor_set(x_166, 6, x_153); +lean_ctor_set(x_166, 7, x_154); +lean_ctor_set(x_166, 8, x_155); +lean_ctor_set(x_166, 9, x_156); +lean_ctor_set(x_166, 10, x_142); +lean_ctor_set(x_166, 11, x_157); +lean_ctor_set(x_166, 12, x_158); +lean_ctor_set(x_166, 13, x_159); +lean_ctor_set(x_166, 14, x_160); +lean_ctor_set(x_166, 15, x_161); +lean_ctor_set(x_166, 16, x_162); +lean_ctor_set(x_166, 17, x_163); +lean_ctor_set(x_166, 18, x_164); +lean_ctor_set_uint8(x_166, sizeof(void*)*19, x_152); +x_167 = lean_st_ref_set(x_3, x_166, x_145); +x_168 = lean_ctor_get(x_167, 1); +lean_inc(x_168); +lean_dec(x_167); +x_169 = lean_st_ref_get(x_3, x_168); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_169, 1); +lean_inc(x_171); +lean_dec(x_169); +x_172 = lean_ctor_get(x_170, 4); +lean_inc(x_172); +lean_dec(x_170); +x_173 = lean_box(0); +x_174 = lean_box(0); +lean_inc(x_140); +x_175 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_2, x_140, x_172, x_173, x_140, x_140, x_174, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_171); +lean_dec(x_140); +if (lean_obj_tag(x_175) == 0) +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_176 = lean_ctor_get(x_175, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_177 = x_175; +} else { + lean_dec_ref(x_175); + x_177 = lean_box(0); +} +if (lean_is_scalar(x_177)) { + x_178 = lean_alloc_ctor(0, 2, 0); +} else { + x_178 = x_177; +} +lean_ctor_set(x_178, 0, x_174); +lean_ctor_set(x_178, 1, x_176); +return x_178; +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_179 = lean_ctor_get(x_175, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_175, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_181 = x_175; +} else { + lean_dec_ref(x_175); + x_181 = lean_box(0); +} +if (lean_is_scalar(x_181)) { + x_182 = lean_alloc_ctor(1, 2, 0); +} else { + x_182 = x_181; +} +lean_ctor_set(x_182, 0, x_179); +lean_ctor_set(x_182, 1, x_180); +return x_182; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_13 = l_List_reverse___rarg(x_3); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_2); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_2, 0); +x_17 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_18 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_16, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_19); +{ +lean_object* _tmp_1 = x_17; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_11 = x_20; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +else +{ +uint8_t x_22; +lean_free_object(x_2); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) +{ +return x_18; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_2, 0); +x_27 = lean_ctor_get(x_2, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_2); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_28 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern(x_26, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_3); +x_2 = x_27; +x_3 = x_31; +x_12 = x_30; +goto _start; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_27); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_33 = lean_ctor_get(x_28, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_35 = x_28; +} else { + lean_dec_ref(x_28); + x_35 = lean_box(0); +} +if (lean_is_scalar(x_35)) { + x_36 = lean_alloc_ctor(1, 2, 0); +} else { + x_36 = x_35; +} +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_34); +return x_36; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_13, 9); +x_17 = l_Lean_PersistentArray_push___rarg(x_16, x_1); +lean_ctor_set(x_13, 9, x_17); +x_18 = lean_st_ref_set(x_3, x_13, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +x_28 = lean_ctor_get(x_13, 3); +x_29 = lean_ctor_get(x_13, 4); +x_30 = lean_ctor_get(x_13, 5); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*19); +x_32 = lean_ctor_get(x_13, 6); +x_33 = lean_ctor_get(x_13, 7); +x_34 = lean_ctor_get(x_13, 8); +x_35 = lean_ctor_get(x_13, 9); +x_36 = lean_ctor_get(x_13, 10); +x_37 = lean_ctor_get(x_13, 11); +x_38 = lean_ctor_get(x_13, 12); +x_39 = lean_ctor_get(x_13, 13); +x_40 = lean_ctor_get(x_13, 14); +x_41 = lean_ctor_get(x_13, 15); +x_42 = lean_ctor_get(x_13, 16); +x_43 = lean_ctor_get(x_13, 17); +x_44 = lean_ctor_get(x_13, 18); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_45 = l_Lean_PersistentArray_push___rarg(x_35, x_1); +x_46 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_46, 0, x_25); +lean_ctor_set(x_46, 1, x_26); +lean_ctor_set(x_46, 2, x_27); +lean_ctor_set(x_46, 3, x_28); +lean_ctor_set(x_46, 4, x_29); +lean_ctor_set(x_46, 5, x_30); +lean_ctor_set(x_46, 6, x_32); +lean_ctor_set(x_46, 7, x_33); +lean_ctor_set(x_46, 8, x_34); +lean_ctor_set(x_46, 9, x_45); +lean_ctor_set(x_46, 10, x_36); +lean_ctor_set(x_46, 11, x_37); +lean_ctor_set(x_46, 12, x_38); +lean_ctor_set(x_46, 13, x_39); +lean_ctor_set(x_46, 14, x_40); +lean_ctor_set(x_46, 15, x_41); +lean_ctor_set(x_46, 16, x_42); +lean_ctor_set(x_46, 17, x_43); +lean_ctor_set(x_46, 18, x_44); +lean_ctor_set_uint8(x_46, sizeof(void*)*19, x_31); +x_47 = lean_st_ref_set(x_3, x_46, x_14); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); +} +x_50 = lean_box(0); +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_49; +} +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("activated `", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`, ", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_1); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_ctor_get(x_1, 1); +x_15 = lean_ctor_get(x_1, 2); +x_16 = lean_ctor_get(x_1, 3); +x_17 = lean_ctor_get(x_1, 4); +x_18 = lean_ctor_get(x_1, 5); +x_19 = l_Lean_Meta_Grind_shareCommon(x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +x_23 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_24 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___spec__1(x_2, x_16, x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +lean_inc(x_18); +lean_inc(x_25); +lean_ctor_set(x_1, 3, x_25); +lean_ctor_set(x_1, 1, x_21); +x_27 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_28 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_26); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_unbox(x_29); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_25); +lean_free_object(x_19); +lean_dec(x_18); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_dec(x_28); +x_32 = lean_box(0); +x_33 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_33; +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_28); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_28, 1); +x_36 = lean_ctor_get(x_28, 0); +lean_dec(x_36); +x_37 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Lean_Meta_Grind_Origin_key(x_18); +lean_dec(x_18); +x_40 = l_Lean_MessageData_ofName(x_39); +x_41 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2; +lean_ctor_set_tag(x_28, 7); +lean_ctor_set(x_28, 1, x_40); +lean_ctor_set(x_28, 0, x_41); +x_42 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4; +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_42); +lean_ctor_set(x_19, 0, x_28); +x_43 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_25, x_23); +x_44 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_43, x_23); +x_45 = l_Lean_MessageData_ofList(x_44); +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_19); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_27, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_38); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_50, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_51); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_50); +return x_52; +} +else +{ +uint8_t x_53; +lean_free_object(x_28); +lean_dec(x_1); +lean_dec(x_25); +lean_free_object(x_19); +lean_dec(x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_53 = !lean_is_exclusive(x_37); +if (x_53 == 0) +{ +return x_37; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_37, 0); +x_55 = lean_ctor_get(x_37, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_37); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; +} +} +} +else +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_28, 1); +lean_inc(x_57); +lean_dec(x_28); +x_58 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_57); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = l_Lean_Meta_Grind_Origin_key(x_18); +lean_dec(x_18); +x_61 = l_Lean_MessageData_ofName(x_60); +x_62 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2; +x_63 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +x_64 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4; +lean_ctor_set_tag(x_19, 7); +lean_ctor_set(x_19, 1, x_64); +lean_ctor_set(x_19, 0, x_63); +x_65 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_25, x_23); +x_66 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_65, x_23); +x_67 = l_Lean_MessageData_ofList(x_66); +x_68 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_68, 0, x_19); +lean_ctor_set(x_68, 1, x_67); +x_69 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_27, x_70, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_59); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +x_74 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_72, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_73); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_72); +return x_74; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_1); +lean_dec(x_25); +lean_free_object(x_19); +lean_dec(x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_75 = lean_ctor_get(x_58, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_58, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_77 = x_58; +} else { + lean_dec_ref(x_58); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +} +else +{ +uint8_t x_79; +lean_free_object(x_19); +lean_dec(x_21); +lean_free_object(x_1); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_79 = !lean_is_exclusive(x_24); +if (x_79 == 0) +{ +return x_24; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_24, 0); +x_81 = lean_ctor_get(x_24, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_24); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; +} +} +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_83 = lean_ctor_get(x_19, 0); +x_84 = lean_ctor_get(x_19, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_19); +x_85 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_86 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___spec__1(x_2, x_16, x_85, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_84); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +lean_inc(x_18); +lean_inc(x_87); +lean_ctor_set(x_1, 3, x_87); +lean_ctor_set(x_1, 1, x_83); +x_89 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_90 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_89, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_88); +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_unbox(x_91); +lean_dec(x_91); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_87); +lean_dec(x_18); +x_93 = lean_ctor_get(x_90, 1); +lean_inc(x_93); +lean_dec(x_90); +x_94 = lean_box(0); +x_95 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_94, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_93); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_95; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_90, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_97 = x_90; +} else { + lean_dec_ref(x_90); + x_97 = lean_box(0); +} +x_98 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_96); +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +lean_dec(x_98); +x_100 = l_Lean_Meta_Grind_Origin_key(x_18); +lean_dec(x_18); +x_101 = l_Lean_MessageData_ofName(x_100); +x_102 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2; +if (lean_is_scalar(x_97)) { + x_103 = lean_alloc_ctor(7, 2, 0); +} else { + x_103 = x_97; + lean_ctor_set_tag(x_103, 7); +} +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_101); +x_104 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4; +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 1, x_104); +x_106 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_87, x_85); +x_107 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_106, x_85); +x_108 = l_Lean_MessageData_ofList(x_107); +x_109 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_109, 0, x_105); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_111 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_89, x_111, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_99); +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +lean_dec(x_112); +x_115 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_113, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_114); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_113); +return x_115; +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_97); +lean_dec(x_1); +lean_dec(x_87); +lean_dec(x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_116 = lean_ctor_get(x_98, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_98, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_118 = x_98; +} else { + lean_dec_ref(x_98); + x_118 = lean_box(0); +} +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); +} else { + x_119 = x_118; +} +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; +} +} +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_83); +lean_free_object(x_1); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_120 = lean_ctor_get(x_86, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_86, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_122 = x_86; +} else { + lean_dec_ref(x_86); + x_122 = lean_box(0); +} +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(1, 2, 0); +} else { + x_123 = x_122; +} +lean_ctor_set(x_123, 0, x_120); +lean_ctor_set(x_123, 1, x_121); +return x_123; +} +} +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_124 = lean_ctor_get(x_1, 0); +x_125 = lean_ctor_get(x_1, 1); +x_126 = lean_ctor_get(x_1, 2); +x_127 = lean_ctor_get(x_1, 3); +x_128 = lean_ctor_get(x_1, 4); +x_129 = lean_ctor_get(x_1, 5); +lean_inc(x_129); +lean_inc(x_128); +lean_inc(x_127); +lean_inc(x_126); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_1); +x_130 = l_Lean_Meta_Grind_shareCommon(x_125, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_133 = x_130; +} else { + lean_dec_ref(x_130); + x_133 = lean_box(0); +} +x_134 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_135 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___spec__1(x_2, x_127, x_134, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_132); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_135, 1); +lean_inc(x_137); +lean_dec(x_135); +lean_inc(x_129); +lean_inc(x_136); +x_138 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_138, 0, x_124); +lean_ctor_set(x_138, 1, x_131); +lean_ctor_set(x_138, 2, x_126); +lean_ctor_set(x_138, 3, x_136); +lean_ctor_set(x_138, 4, x_128); +lean_ctor_set(x_138, 5, x_129); +x_139 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2; +x_140 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_139, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_137); +x_141 = lean_ctor_get(x_140, 0); +lean_inc(x_141); +x_142 = lean_unbox(x_141); +lean_dec(x_141); +if (x_142 == 0) +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_136); +lean_dec(x_133); +lean_dec(x_129); +x_143 = lean_ctor_get(x_140, 1); +lean_inc(x_143); +lean_dec(x_140); +x_144 = lean_box(0); +x_145 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(x_138, x_144, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_143); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_145; +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_140, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_147 = x_140; +} else { + lean_dec_ref(x_140); + x_147 = lean_box(0); +} +x_148 = l_Lean_Meta_Grind_updateLastTag(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_146); +if (lean_obj_tag(x_148) == 0) +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_149 = lean_ctor_get(x_148, 1); +lean_inc(x_149); +lean_dec(x_148); +x_150 = l_Lean_Meta_Grind_Origin_key(x_129); +lean_dec(x_129); +x_151 = l_Lean_MessageData_ofName(x_150); +x_152 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2; +if (lean_is_scalar(x_147)) { + x_153 = lean_alloc_ctor(7, 2, 0); +} else { + x_153 = x_147; + lean_ctor_set_tag(x_153, 7); +} +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_151); +x_154 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4; +if (lean_is_scalar(x_133)) { + x_155 = lean_alloc_ctor(7, 2, 0); +} else { + x_155 = x_133; + lean_ctor_set_tag(x_155, 7); +} +lean_ctor_set(x_155, 0, x_153); +lean_ctor_set(x_155, 1, x_154); +x_156 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__1(x_136, x_134); +x_157 = l_List_mapTR_loop___at_Lean_Meta_Grind_mkEMatchTheoremCore___spec__2(x_156, x_134); +x_158 = l_Lean_MessageData_ofList(x_157); +x_159 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_159, 0, x_155); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6; +x_161 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +x_162 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_139, x_161, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_149); +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(x_138, x_163, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_164); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_163); +return x_165; +} +else +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +lean_dec(x_147); +lean_dec(x_138); +lean_dec(x_136); +lean_dec(x_133); +lean_dec(x_129); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_166 = lean_ctor_get(x_148, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_148, 1); +lean_inc(x_167); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_168 = x_148; } else { - lean_dec_ref(x_93); - x_126 = lean_box(0); + lean_dec_ref(x_148); + x_168 = lean_box(0); } -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_168)) { + x_169 = lean_alloc_ctor(1, 2, 0); } else { - x_127 = x_126; + x_169 = x_168; } -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_ctor_set(x_169, 0, x_166); +lean_ctor_set(x_169, 1, x_167); +return x_169; } } } else { -uint8_t x_128; -lean_free_object(x_6); -lean_inc(x_33); -lean_inc(x_35); -lean_ctor_set(x_20, 4, x_35); -x_128 = !lean_is_exclusive(x_35); -if (x_128 == 0) -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_129 = lean_ctor_get(x_35, 1); -lean_dec(x_129); -x_130 = lean_ctor_get(x_35, 0); -lean_dec(x_130); -x_131 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_132 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_131, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_unbox(x_133); +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_dec(x_133); -if (x_134 == 0) -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; -lean_free_object(x_35); -lean_dec(x_33); -x_135 = lean_ctor_get(x_132, 1); -lean_inc(x_135); -lean_dec(x_132); -x_136 = lean_box(0); -x_137 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_136, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_135); -x_138 = lean_ctor_get(x_137, 0); -lean_inc(x_138); -x_139 = lean_ctor_get(x_137, 1); -lean_inc(x_139); -lean_dec(x_137); -x_22 = x_138; -x_23 = x_139; -goto block_26; +lean_dec(x_131); +lean_dec(x_129); +lean_dec(x_128); +lean_dec(x_126); +lean_dec(x_124); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_170 = lean_ctor_get(x_135, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_135, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_135)) { + lean_ctor_release(x_135, 0); + lean_ctor_release(x_135, 1); + x_172 = x_135; +} else { + lean_dec_ref(x_135); + x_172 = lean_box(0); } -else -{ -uint8_t x_140; -x_140 = !lean_is_exclusive(x_132); -if (x_140 == 0) -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_141 = lean_ctor_get(x_132, 1); -x_142 = lean_ctor_get(x_132, 0); -lean_dec(x_142); -x_143 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_144 = l_Lean_MessageData_ofName(x_143); -x_145 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -lean_ctor_set_tag(x_132, 7); -lean_ctor_set(x_132, 1, x_144); -lean_ctor_set(x_132, 0, x_145); -x_146 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -lean_ctor_set_tag(x_35, 7); -lean_ctor_set(x_35, 1, x_146); -lean_ctor_set(x_35, 0, x_132); -x_147 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_131, x_35, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_141); -x_148 = lean_ctor_get(x_147, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_147, 1); -lean_inc(x_149); -lean_dec(x_147); -x_150 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_148, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_149); -lean_dec(x_148); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec(x_150); -x_22 = x_151; -x_23 = x_152; -goto block_26; +if (lean_is_scalar(x_172)) { + x_173 = lean_alloc_ctor(1, 2, 0); +} else { + x_173 = x_172; } -else -{ -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_153 = lean_ctor_get(x_132, 1); -lean_inc(x_153); -lean_dec(x_132); -x_154 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_155 = l_Lean_MessageData_ofName(x_154); -x_156 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -x_157 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_155); -x_158 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -lean_ctor_set_tag(x_35, 7); -lean_ctor_set(x_35, 1, x_158); -lean_ctor_set(x_35, 0, x_157); -x_159 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_131, x_35, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_153); -x_160 = lean_ctor_get(x_159, 0); -lean_inc(x_160); -x_161 = lean_ctor_get(x_159, 1); -lean_inc(x_161); -lean_dec(x_159); -x_162 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_160, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_161); -lean_dec(x_160); -x_163 = lean_ctor_get(x_162, 0); -lean_inc(x_163); -x_164 = lean_ctor_get(x_162, 1); -lean_inc(x_164); -lean_dec(x_162); -x_22 = x_163; -x_23 = x_164; -goto block_26; +lean_ctor_set(x_173, 0, x_170); +lean_ctor_set(x_173, 1, x_171); +return x_173; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; -lean_dec(x_35); -x_165 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_166 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_165, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_167 = lean_ctor_get(x_166, 0); -lean_inc(x_167); -x_168 = lean_unbox(x_167); -lean_dec(x_167); -if (x_168 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -lean_dec(x_33); -x_169 = lean_ctor_get(x_166, 1); -lean_inc(x_169); -lean_dec(x_166); -x_170 = lean_box(0); -x_171 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_170, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_169); -x_172 = lean_ctor_get(x_171, 0); -lean_inc(x_172); -x_173 = lean_ctor_get(x_171, 1); -lean_inc(x_173); -lean_dec(x_171); -x_22 = x_172; -x_23 = x_173; -goto block_26; +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_is_matcher(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -x_174 = lean_ctor_get(x_166, 1); -lean_inc(x_174); -if (lean_is_exclusive(x_166)) { - lean_ctor_release(x_166, 0); - lean_ctor_release(x_166, 1); - x_175 = x_166; -} else { - lean_dec_ref(x_166); - x_175 = lean_box(0); -} -x_176 = l_Lean_Meta_Grind_Origin_key(x_33); -lean_dec(x_33); -x_177 = l_Lean_MessageData_ofName(x_176); -x_178 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -if (lean_is_scalar(x_175)) { - x_179 = lean_alloc_ctor(7, 2, 0); -} else { - x_179 = x_175; - lean_ctor_set_tag(x_179, 7); -} -lean_ctor_set(x_179, 0, x_178); -lean_ctor_set(x_179, 1, x_177); -x_180 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -x_181 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_181, 0, x_179); -lean_ctor_set(x_181, 1, x_180); -x_182 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_165, x_181, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_174); -x_183 = lean_ctor_get(x_182, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_182, 1); -lean_inc(x_184); -lean_dec(x_182); -x_185 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_20, x_183, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_184); -lean_dec(x_183); -x_186 = lean_ctor_get(x_185, 0); -lean_inc(x_186); -x_187 = lean_ctor_get(x_185, 1); -lean_inc(x_187); -lean_dec(x_185); -x_22 = x_186; -x_23 = x_187; -goto block_26; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_is_matcher(x_19, x_1); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; } } } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_6, x_5); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_16); +return x_18; } else { -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_188 = lean_ctor_get(x_20, 0); -x_189 = lean_ctor_get(x_20, 1); -x_190 = lean_ctor_get(x_20, 2); -x_191 = lean_ctor_get(x_20, 3); -x_192 = lean_ctor_get(x_20, 4); -x_193 = lean_ctor_get(x_20, 5); -lean_inc(x_193); -lean_inc(x_192); -lean_inc(x_191); -lean_inc(x_190); -lean_inc(x_189); -lean_inc(x_188); -lean_dec(x_20); -x_194 = lean_box(0); -lean_inc(x_3); -x_195 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_192, x_194); -if (lean_obj_tag(x_195) == 0) +lean_object* x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; +lean_dec(x_7); +x_19 = lean_array_uget(x_4, x_6); +x_20 = 0; +x_21 = 1; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_22 = l_Lean_Meta_Grind_mkEMatchEqTheorem(x_19, x_20, x_21, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_196 = l_Lean_Meta_Grind_shareCommon(x_189, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_197 = lean_ctor_get(x_196, 0); -lean_inc(x_197); -x_198 = lean_ctor_get(x_196, 1); -lean_inc(x_198); -if (lean_is_exclusive(x_196)) { - lean_ctor_release(x_196, 0); - lean_ctor_release(x_196, 1); - x_199 = x_196; -} else { - lean_dec_ref(x_196); - x_199 = lean_box(0); -} -lean_inc(x_16); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); @@ -10127,125 +14766,26 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); +lean_inc(x_8); lean_inc(x_1); -x_200 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_191, x_194, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_198); -if (lean_obj_tag(x_200) == 0) -{ -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; uint8_t x_207; -x_201 = lean_ctor_get(x_200, 0); -lean_inc(x_201); -x_202 = lean_ctor_get(x_200, 1); -lean_inc(x_202); -lean_dec(x_200); -lean_inc(x_193); -lean_inc(x_201); -x_203 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_203, 0, x_188); -lean_ctor_set(x_203, 1, x_197); -lean_ctor_set(x_203, 2, x_190); -lean_ctor_set(x_203, 3, x_201); -lean_ctor_set(x_203, 4, x_195); -lean_ctor_set(x_203, 5, x_193); -x_204 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_205 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_204, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_202); -x_206 = lean_ctor_get(x_205, 0); -lean_inc(x_206); -x_207 = lean_unbox(x_206); -lean_dec(x_206); -if (x_207 == 0) +x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem(x_23, x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; -lean_dec(x_201); -lean_dec(x_199); -lean_dec(x_193); -lean_free_object(x_6); -x_208 = lean_ctor_get(x_205, 1); -lean_inc(x_208); -lean_dec(x_205); -x_209 = lean_box(0); -x_210 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_203, x_209, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_208); -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_210, 1); -lean_inc(x_212); -lean_dec(x_210); -x_22 = x_211; -x_23 = x_212; -goto block_26; -} -else -{ -lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_213 = lean_ctor_get(x_205, 1); -lean_inc(x_213); -if (lean_is_exclusive(x_205)) { - lean_ctor_release(x_205, 0); - lean_ctor_release(x_205, 1); - x_214 = x_205; -} else { - lean_dec_ref(x_205); - x_214 = lean_box(0); -} -x_215 = l_Lean_Meta_Grind_Origin_key(x_193); -lean_dec(x_193); -x_216 = l_Lean_MessageData_ofName(x_215); -x_217 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -if (lean_is_scalar(x_214)) { - x_218 = lean_alloc_ctor(7, 2, 0); -} else { - x_218 = x_214; - lean_ctor_set_tag(x_218, 7); -} -lean_ctor_set(x_218, 0, x_217); -lean_ctor_set(x_218, 1, x_216); -x_219 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -if (lean_is_scalar(x_199)) { - x_220 = lean_alloc_ctor(7, 2, 0); -} else { - x_220 = x_199; - lean_ctor_set_tag(x_220, 7); -} -lean_ctor_set(x_220, 0, x_218); -lean_ctor_set(x_220, 1, x_219); -x_221 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_201, x_194); -x_222 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_221, x_194); -x_223 = l_Lean_MessageData_ofList(x_222); -lean_ctor_set_tag(x_6, 7); -lean_ctor_set(x_6, 1, x_223); -lean_ctor_set(x_6, 0, x_220); -x_224 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_225 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_225, 0, x_6); -lean_ctor_set(x_225, 1, x_224); -x_226 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_204, x_225, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_213); -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_226, 1); -lean_inc(x_228); -lean_dec(x_226); -x_229 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_203, x_227, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_228); -lean_dec(x_227); -x_230 = lean_ctor_get(x_229, 0); -lean_inc(x_230); -x_231 = lean_ctor_get(x_229, 1); -lean_inc(x_231); -lean_dec(x_229); -x_22 = x_230; -x_23 = x_231; -goto block_26; -} +lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = 1; +x_28 = lean_usize_add(x_6, x_27); +x_29 = lean_box(0); +x_6 = x_28; +x_7 = x_29; +x_16 = x_26; +goto _start; } else { -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; -lean_dec(x_199); -lean_dec(x_197); -lean_dec(x_193); -lean_dec(x_190); -lean_dec(x_188); -lean_free_object(x_6); -lean_dec(x_21); -lean_dec(x_16); +uint8_t x_31; lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -10253,489 +14793,268 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_3); +lean_dec(x_8); lean_dec(x_1); -x_232 = lean_ctor_get(x_200, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_200, 1); -lean_inc(x_233); -if (lean_is_exclusive(x_200)) { - lean_ctor_release(x_200, 0); - lean_ctor_release(x_200, 1); - x_234 = x_200; -} else { - lean_dec_ref(x_200); - x_234 = lean_box(0); -} -if (lean_is_scalar(x_234)) { - x_235 = lean_alloc_ctor(1, 2, 0); -} else { - x_235 = x_234; -} -lean_ctor_set(x_235, 0, x_232); -lean_ctor_set(x_235, 1, x_233); -return x_235; -} +x_31 = !lean_is_exclusive(x_25); +if (x_31 == 0) +{ +return x_25; } else { -lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; uint8_t x_241; -lean_free_object(x_6); -lean_inc(x_193); -lean_inc(x_195); -x_236 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_236, 0, x_188); -lean_ctor_set(x_236, 1, x_189); -lean_ctor_set(x_236, 2, x_190); -lean_ctor_set(x_236, 3, x_191); -lean_ctor_set(x_236, 4, x_195); -lean_ctor_set(x_236, 5, x_193); -if (lean_is_exclusive(x_195)) { - lean_ctor_release(x_195, 0); - lean_ctor_release(x_195, 1); - x_237 = x_195; -} else { - lean_dec_ref(x_195); - x_237 = lean_box(0); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_25, 0); +x_33 = lean_ctor_get(x_25, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_25); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } -x_238 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_239 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_238, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_240 = lean_ctor_get(x_239, 0); -lean_inc(x_240); -x_241 = lean_unbox(x_240); -lean_dec(x_240); -if (x_241 == 0) -{ -lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -lean_dec(x_237); -lean_dec(x_193); -x_242 = lean_ctor_get(x_239, 1); -lean_inc(x_242); -lean_dec(x_239); -x_243 = lean_box(0); -x_244 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_236, x_243, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_242); -x_245 = lean_ctor_get(x_244, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_244, 1); -lean_inc(x_246); -lean_dec(x_244); -x_22 = x_245; -x_23 = x_246; -goto block_26; -} -else -{ -lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; -x_247 = lean_ctor_get(x_239, 1); -lean_inc(x_247); -if (lean_is_exclusive(x_239)) { - lean_ctor_release(x_239, 0); - lean_ctor_release(x_239, 1); - x_248 = x_239; -} else { - lean_dec_ref(x_239); - x_248 = lean_box(0); -} -x_249 = l_Lean_Meta_Grind_Origin_key(x_193); -lean_dec(x_193); -x_250 = l_Lean_MessageData_ofName(x_249); -x_251 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -if (lean_is_scalar(x_248)) { - x_252 = lean_alloc_ctor(7, 2, 0); -} else { - x_252 = x_248; - lean_ctor_set_tag(x_252, 7); -} -lean_ctor_set(x_252, 0, x_251); -lean_ctor_set(x_252, 1, x_250); -x_253 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -if (lean_is_scalar(x_237)) { - x_254 = lean_alloc_ctor(7, 2, 0); -} else { - x_254 = x_237; - lean_ctor_set_tag(x_254, 7); -} -lean_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_253); -x_255 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_238, x_254, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_247); -x_256 = lean_ctor_get(x_255, 0); -lean_inc(x_256); -x_257 = lean_ctor_get(x_255, 1); -lean_inc(x_257); -lean_dec(x_255); -x_258 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_236, x_256, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_257); -lean_dec(x_256); -x_259 = lean_ctor_get(x_258, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_258, 1); -lean_inc(x_260); -lean_dec(x_258); -x_22 = x_259; -x_23 = x_260; -goto block_26; } } +else +{ +uint8_t x_35; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_22); +if (x_35 == 0) +{ +return x_22; } -block_26: +else { -lean_object* x_24; -x_24 = lean_ctor_get(x_22, 0); -lean_inc(x_24); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_22, 0); +x_37 = lean_ctor_get(x_22, 1); +lean_inc(x_37); +lean_inc(x_36); lean_dec(x_22); -x_6 = x_21; -x_7 = x_24; -x_8 = lean_box(0); -x_17 = x_23; -goto _start; +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } -else -{ -lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; -x_261 = lean_ctor_get(x_6, 0); -x_262 = lean_ctor_get(x_6, 1); -lean_inc(x_262); -lean_inc(x_261); -lean_dec(x_6); -x_268 = lean_ctor_get(x_261, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_261, 1); -lean_inc(x_269); -x_270 = lean_ctor_get(x_261, 2); -lean_inc(x_270); -x_271 = lean_ctor_get(x_261, 3); -lean_inc(x_271); -x_272 = lean_ctor_get(x_261, 4); -lean_inc(x_272); -x_273 = lean_ctor_get(x_261, 5); -lean_inc(x_273); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - lean_ctor_release(x_261, 2); - lean_ctor_release(x_261, 3); - lean_ctor_release(x_261, 4); - lean_ctor_release(x_261, 5); - x_274 = x_261; -} else { - lean_dec_ref(x_261); - x_274 = lean_box(0); } -x_275 = lean_box(0); -lean_inc(x_3); -x_276 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__6(x_3, x_272, x_275); -if (lean_obj_tag(x_276) == 0) -{ -lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -x_277 = l_Lean_Meta_Grind_shareCommon(x_269, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_278 = lean_ctor_get(x_277, 0); -lean_inc(x_278); -x_279 = lean_ctor_get(x_277, 1); -lean_inc(x_279); -if (lean_is_exclusive(x_277)) { - lean_ctor_release(x_277, 0); - lean_ctor_release(x_277, 1); - x_280 = x_277; -} else { - lean_dec_ref(x_277); - x_280 = lean_box(0); } -lean_inc(x_16); -lean_inc(x_15); +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_13 = lean_st_ref_take(x_4, x_12); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +x_18 = lean_ctor_get(x_14, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_14, 3); +lean_inc(x_19); +x_20 = lean_ctor_get(x_14, 4); +lean_inc(x_20); +x_21 = lean_ctor_get(x_14, 5); +lean_inc(x_21); +x_22 = lean_ctor_get_uint8(x_14, sizeof(void*)*19); +x_23 = lean_ctor_get(x_14, 6); +lean_inc(x_23); +x_24 = lean_ctor_get(x_14, 7); +lean_inc(x_24); +x_25 = lean_ctor_get(x_14, 8); +lean_inc(x_25); +x_26 = lean_ctor_get(x_14, 9); +lean_inc(x_26); +x_27 = lean_ctor_get(x_14, 10); +lean_inc(x_27); +x_28 = lean_ctor_get(x_14, 11); +lean_inc(x_28); +x_29 = lean_ctor_get(x_14, 12); +lean_inc(x_29); +x_30 = lean_ctor_get(x_14, 13); +lean_inc(x_30); +x_31 = lean_ctor_get(x_14, 14); +lean_inc(x_31); +x_32 = lean_ctor_get(x_14, 15); +lean_inc(x_32); +x_33 = lean_box(0); +lean_inc(x_1); +x_34 = l_Lean_PersistentHashMap_insert___at_Lean_NameSSet_insert___spec__2(x_32, x_1, x_33); +x_35 = lean_ctor_get(x_14, 16); +lean_inc(x_35); +x_36 = lean_ctor_get(x_14, 17); +lean_inc(x_36); +x_37 = lean_ctor_get(x_14, 18); +lean_inc(x_37); +lean_dec(x_14); +x_38 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_38, 0, x_16); +lean_ctor_set(x_38, 1, x_17); +lean_ctor_set(x_38, 2, x_18); +lean_ctor_set(x_38, 3, x_19); +lean_ctor_set(x_38, 4, x_20); +lean_ctor_set(x_38, 5, x_21); +lean_ctor_set(x_38, 6, x_23); +lean_ctor_set(x_38, 7, x_24); +lean_ctor_set(x_38, 8, x_25); +lean_ctor_set(x_38, 9, x_26); +lean_ctor_set(x_38, 10, x_27); +lean_ctor_set(x_38, 11, x_28); +lean_ctor_set(x_38, 12, x_29); +lean_ctor_set(x_38, 13, x_30); +lean_ctor_set(x_38, 14, x_31); +lean_ctor_set(x_38, 15, x_34); +lean_ctor_set(x_38, 16, x_35); +lean_ctor_set(x_38, 17, x_36); +lean_ctor_set(x_38, 18, x_37); +lean_ctor_set_uint8(x_38, sizeof(void*)*19, x_22); +x_39 = lean_st_ref_set(x_4, x_38, x_15); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_1); -x_281 = l_List_mapM_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__7(x_1, x_271, x_275, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_279); -if (lean_obj_tag(x_281) == 0) +lean_inc(x_8); +x_41 = lean_get_match_equations_for(x_1, x_8, x_9, x_10, x_11, x_40); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; uint8_t x_288; -x_282 = lean_ctor_get(x_281, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_281, 1); -lean_inc(x_283); -lean_dec(x_281); -lean_inc(x_273); -lean_inc(x_282); -if (lean_is_scalar(x_274)) { - x_284 = lean_alloc_ctor(0, 6, 0); -} else { - x_284 = x_274; -} -lean_ctor_set(x_284, 0, x_268); -lean_ctor_set(x_284, 1, x_278); -lean_ctor_set(x_284, 2, x_270); -lean_ctor_set(x_284, 3, x_282); -lean_ctor_set(x_284, 4, x_276); -lean_ctor_set(x_284, 5, x_273); -x_285 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_286 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_285, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_283); -x_287 = lean_ctor_get(x_286, 0); -lean_inc(x_287); -x_288 = lean_unbox(x_287); -lean_dec(x_287); -if (x_288 == 0) +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; size_t x_46; size_t x_47; lean_object* x_48; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_box(0); +x_46 = lean_array_size(x_44); +x_47 = 0; +x_48 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2(x_2, x_44, x_45, x_44, x_46, x_47, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_43); +lean_dec(x_44); +if (lean_obj_tag(x_48) == 0) { -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; -lean_dec(x_282); -lean_dec(x_280); -lean_dec(x_273); -x_289 = lean_ctor_get(x_286, 1); -lean_inc(x_289); -lean_dec(x_286); -x_290 = lean_box(0); -x_291 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_284, x_290, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_289); -x_292 = lean_ctor_get(x_291, 0); -lean_inc(x_292); -x_293 = lean_ctor_get(x_291, 1); -lean_inc(x_293); -lean_dec(x_291); -x_263 = x_292; -x_264 = x_293; -goto block_267; +uint8_t x_49; +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) +{ +lean_object* x_50; +x_50 = lean_ctor_get(x_48, 0); +lean_dec(x_50); +lean_ctor_set(x_48, 0, x_33); +return x_48; } else { -lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; -x_294 = lean_ctor_get(x_286, 1); -lean_inc(x_294); -if (lean_is_exclusive(x_286)) { - lean_ctor_release(x_286, 0); - lean_ctor_release(x_286, 1); - x_295 = x_286; -} else { - lean_dec_ref(x_286); - x_295 = lean_box(0); -} -x_296 = l_Lean_Meta_Grind_Origin_key(x_273); -lean_dec(x_273); -x_297 = l_Lean_MessageData_ofName(x_296); -x_298 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4; -if (lean_is_scalar(x_295)) { - x_299 = lean_alloc_ctor(7, 2, 0); -} else { - x_299 = x_295; - lean_ctor_set_tag(x_299, 7); -} -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_299, 1, x_297); -x_300 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6; -if (lean_is_scalar(x_280)) { - x_301 = lean_alloc_ctor(7, 2, 0); -} else { - x_301 = x_280; - lean_ctor_set_tag(x_301, 7); -} -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_300); -x_302 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__3(x_282, x_275); -x_303 = l_List_mapTR_loop___at_Lean_Meta_Grind_addEMatchTheorem___spec__4(x_302, x_275); -x_304 = l_Lean_MessageData_ofList(x_303); -x_305 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_305, 0, x_301); -lean_ctor_set(x_305, 1, x_304); -x_306 = l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__5; -x_307 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_307, 0, x_305); -lean_ctor_set(x_307, 1, x_306); -x_308 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_285, x_307, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_294); -x_309 = lean_ctor_get(x_308, 0); -lean_inc(x_309); -x_310 = lean_ctor_get(x_308, 1); -lean_inc(x_310); -lean_dec(x_308); -x_311 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_284, x_309, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_310); -lean_dec(x_309); -x_312 = lean_ctor_get(x_311, 0); -lean_inc(x_312); -x_313 = lean_ctor_get(x_311, 1); -lean_inc(x_313); -lean_dec(x_311); -x_263 = x_312; -x_264 = x_313; -goto block_267; +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_48, 1); +lean_inc(x_51); +lean_dec(x_48); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_33); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } else { -lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; -lean_dec(x_280); -lean_dec(x_278); -lean_dec(x_274); -lean_dec(x_273); -lean_dec(x_270); -lean_dec(x_268); -lean_dec(x_262); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_3); -lean_dec(x_1); -x_314 = lean_ctor_get(x_281, 0); -lean_inc(x_314); -x_315 = lean_ctor_get(x_281, 1); -lean_inc(x_315); -if (lean_is_exclusive(x_281)) { - lean_ctor_release(x_281, 0); - lean_ctor_release(x_281, 1); - x_316 = x_281; -} else { - lean_dec_ref(x_281); - x_316 = lean_box(0); +uint8_t x_53; +x_53 = !lean_is_exclusive(x_48); +if (x_53 == 0) +{ +return x_48; } -if (lean_is_scalar(x_316)) { - x_317 = lean_alloc_ctor(1, 2, 0); -} else { - x_317 = x_316; +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_48, 0); +x_55 = lean_ctor_get(x_48, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_48); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } -lean_ctor_set(x_317, 0, x_314); -lean_ctor_set(x_317, 1, x_315); -return x_317; } } else { -lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; uint8_t x_323; -lean_inc(x_273); -lean_inc(x_276); -if (lean_is_scalar(x_274)) { - x_318 = lean_alloc_ctor(0, 6, 0); -} else { - x_318 = x_274; -} -lean_ctor_set(x_318, 0, x_268); -lean_ctor_set(x_318, 1, x_269); -lean_ctor_set(x_318, 2, x_270); -lean_ctor_set(x_318, 3, x_271); -lean_ctor_set(x_318, 4, x_276); -lean_ctor_set(x_318, 5, x_273); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - lean_ctor_release(x_276, 1); - x_319 = x_276; -} else { - lean_dec_ref(x_276); - x_319 = lean_box(0); -} -x_320 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2; -x_321 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_320, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_322 = lean_ctor_get(x_321, 0); -lean_inc(x_322); -x_323 = lean_unbox(x_322); -lean_dec(x_322); -if (x_323 == 0) +uint8_t x_57; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_57 = !lean_is_exclusive(x_41); +if (x_57 == 0) { -lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; -lean_dec(x_319); -lean_dec(x_273); -x_324 = lean_ctor_get(x_321, 1); -lean_inc(x_324); -lean_dec(x_321); -x_325 = lean_box(0); -x_326 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_318, x_325, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_324); -x_327 = lean_ctor_get(x_326, 0); -lean_inc(x_327); -x_328 = lean_ctor_get(x_326, 1); -lean_inc(x_328); -lean_dec(x_326); -x_263 = x_327; -x_264 = x_328; -goto block_267; +return x_41; } else { -lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; -x_329 = lean_ctor_get(x_321, 1); -lean_inc(x_329); -if (lean_is_exclusive(x_321)) { - lean_ctor_release(x_321, 0); - lean_ctor_release(x_321, 1); - x_330 = x_321; -} else { - lean_dec_ref(x_321); - x_330 = lean_box(0); -} -x_331 = l_Lean_Meta_Grind_Origin_key(x_273); -lean_dec(x_273); -x_332 = l_Lean_MessageData_ofName(x_331); -x_333 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8; -if (lean_is_scalar(x_330)) { - x_334 = lean_alloc_ctor(7, 2, 0); -} else { - x_334 = x_330; - lean_ctor_set_tag(x_334, 7); -} -lean_ctor_set(x_334, 0, x_333); -lean_ctor_set(x_334, 1, x_332); -x_335 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10; -if (lean_is_scalar(x_319)) { - x_336 = lean_alloc_ctor(7, 2, 0); -} else { - x_336 = x_319; - lean_ctor_set_tag(x_336, 7); -} -lean_ctor_set(x_336, 0, x_334); -lean_ctor_set(x_336, 1, x_335); -x_337 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_320, x_336, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_329); -x_338 = lean_ctor_get(x_337, 0); -lean_inc(x_338); -x_339 = lean_ctor_get(x_337, 1); -lean_inc(x_339); -lean_dec(x_337); -x_340 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_318, x_338, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_339); -lean_dec(x_338); -x_341 = lean_ctor_get(x_340, 0); -lean_inc(x_341); -x_342 = lean_ctor_get(x_340, 1); -lean_inc(x_342); -lean_dec(x_340); -x_263 = x_341; -x_264 = x_342; -goto block_267; -} -} -block_267: -{ -lean_object* x_265; -x_265 = lean_ctor_get(x_263, 0); -lean_inc(x_265); -lean_dec(x_263); -x_6 = x_262; -x_7 = x_265; -x_8 = lean_box(0); -x_17 = x_264; -goto _start; -} +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_41, 0); +x_59 = lean_ctor_get(x_41, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_41); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; uint8_t x_13; -x_12 = lean_st_ref_get(x_3, x_11); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_st_ref_get(x_4, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -x_16 = lean_ctor_get(x_14, 10); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_16, x_1); -if (lean_obj_tag(x_17) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_15, 15); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__2(x_17, x_1); +if (x_18 == 0) { -lean_object* x_18; +lean_object* x_19; lean_object* x_20; +lean_free_object(x_13); +x_19 = lean_box(0); +x_20 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1(x_1, x_2, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +return x_20; +} +else +{ +lean_object* x_21; +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10743,229 +15062,149 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -x_18 = lean_box(0); -lean_ctor_set(x_12, 0, x_18); -return x_12; +lean_dec(x_1); +x_21 = lean_box(0); +lean_ctor_set(x_13, 0, x_21); +return x_13; +} } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -lean_free_object(x_12); -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_st_ref_take(x_3, x_15); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_13, 0); +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); lean_inc(x_22); -lean_dec(x_20); -x_23 = !lean_is_exclusive(x_21); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_24 = lean_ctor_get(x_21, 10); -x_25 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_24, x_1); -lean_ctor_set(x_21, 10, x_25); -x_26 = lean_st_ref_set(x_3, x_21, x_22); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_st_ref_get(x_3, x_27); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_ctor_get(x_29, 4); -lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_box(0); -x_33 = lean_box(0); -lean_inc(x_19); -x_34 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_19, x_31, x_32, x_19, x_19, x_33, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -lean_dec(x_19); -if (lean_obj_tag(x_34) == 0) -{ -uint8_t x_35; -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) -{ -lean_object* x_36; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -lean_ctor_set(x_34, 0, x_33); -return x_34; -} -else +lean_dec(x_13); +x_24 = lean_ctor_get(x_22, 15); +lean_inc(x_24); +lean_dec(x_22); +x_25 = l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__2(x_24, x_1); +if (x_25 == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); -lean_dec(x_34); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_33); -lean_ctor_set(x_38, 1, x_37); -return x_38; -} +lean_object* x_26; lean_object* x_27; +x_26 = lean_box(0); +x_27 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1(x_1, x_2, x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_27; } else { -uint8_t x_39; -x_39 = !lean_is_exclusive(x_34); -if (x_39 == 0) -{ -return x_34; +lean_object* x_28; lean_object* x_29; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_23); +return x_29; } -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_34, 0); -x_41 = lean_ctor_get(x_34, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_34); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_43 = lean_ctor_get(x_21, 0); -x_44 = lean_ctor_get(x_21, 1); -x_45 = lean_ctor_get(x_21, 2); -x_46 = lean_ctor_get(x_21, 3); -x_47 = lean_ctor_get(x_21, 4); -x_48 = lean_ctor_get(x_21, 5); -x_49 = lean_ctor_get_uint8(x_21, sizeof(void*)*14); -x_50 = lean_ctor_get(x_21, 6); -x_51 = lean_ctor_get(x_21, 7); -x_52 = lean_ctor_get(x_21, 8); -x_53 = lean_ctor_get(x_21, 9); -x_54 = lean_ctor_get(x_21, 10); -x_55 = lean_ctor_get(x_21, 11); -x_56 = lean_ctor_get(x_21, 12); -x_57 = lean_ctor_get(x_21, 13); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_21); -x_58 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_54, x_1); -x_59 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_59, 0, x_43); -lean_ctor_set(x_59, 1, x_44); -lean_ctor_set(x_59, 2, x_45); -lean_ctor_set(x_59, 3, x_46); -lean_ctor_set(x_59, 4, x_47); -lean_ctor_set(x_59, 5, x_48); -lean_ctor_set(x_59, 6, x_50); -lean_ctor_set(x_59, 7, x_51); -lean_ctor_set(x_59, 8, x_52); -lean_ctor_set(x_59, 9, x_53); -lean_ctor_set(x_59, 10, x_58); -lean_ctor_set(x_59, 11, x_55); -lean_ctor_set(x_59, 12, x_56); -lean_ctor_set(x_59, 13, x_57); -lean_ctor_set_uint8(x_59, sizeof(void*)*14, x_49); -x_60 = lean_st_ref_set(x_3, x_59, x_22); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_62 = lean_st_ref_get(x_3, x_61); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_ctor_get(x_63, 4); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_box(0); -x_67 = lean_box(0); -lean_inc(x_19); -x_68 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_19, x_65, x_66, x_19, x_19, x_67, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_64); -lean_dec(x_19); -if (lean_obj_tag(x_68) == 0) +if (lean_obj_tag(x_1) == 4) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_70 = x_68; -} else { - lean_dec_ref(x_68); - x_70 = lean_box(0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +lean_inc(x_13); +x_14 = l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) +{ +uint8_t x_17; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_17 = !lean_is_exclusive(x_14); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_14, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_14, 0, x_19); +return x_14; } -if (lean_is_scalar(x_70)) { - x_71 = lean_alloc_ctor(0, 2, 0); -} else { - x_71 = x_70; +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_dec(x_14); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } -lean_ctor_set(x_71, 0, x_67); -lean_ctor_set(x_71, 1, x_69); -return x_71; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_68, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_68, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_74 = x_68; -} else { - lean_dec_ref(x_68); - x_74 = lean_box(0); -} -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(1, 2, 0); -} else { - x_75 = x_74; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 1); +lean_inc(x_23); +lean_dec(x_14); +x_24 = lean_box(0); +x_25 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2(x_13, x_2, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_25; } -lean_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; } +else +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_12); +return x_27; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_12, 0); -x_77 = lean_ctor_get(x_12, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_12); -x_78 = lean_ctor_get(x_76, 10); -lean_inc(x_78); -lean_dec(x_76); -x_79 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_EMatchTheorems_insert___spec__2(x_78, x_1); -if (lean_obj_tag(x_79) == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_getConfig___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*4); +lean_dec(x_13); +if (x_14 == 0) { -lean_object* x_80; lean_object* x_81; +uint8_t x_15; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10975,159 +15214,39 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_80 = lean_box(0); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_77); -return x_81; -} -else +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_82 = lean_ctor_get(x_79, 0); -lean_inc(x_82); -lean_dec(x_79); -x_83 = lean_st_ref_take(x_3, x_77); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); -lean_inc(x_85); -lean_dec(x_83); -x_86 = lean_ctor_get(x_84, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_84, 1); -lean_inc(x_87); -x_88 = lean_ctor_get(x_84, 2); -lean_inc(x_88); -x_89 = lean_ctor_get(x_84, 3); -lean_inc(x_89); -x_90 = lean_ctor_get(x_84, 4); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 5); -lean_inc(x_91); -x_92 = lean_ctor_get_uint8(x_84, sizeof(void*)*14); -x_93 = lean_ctor_get(x_84, 6); -lean_inc(x_93); -x_94 = lean_ctor_get(x_84, 7); -lean_inc(x_94); -x_95 = lean_ctor_get(x_84, 8); -lean_inc(x_95); -x_96 = lean_ctor_get(x_84, 9); -lean_inc(x_96); -x_97 = lean_ctor_get(x_84, 10); -lean_inc(x_97); -x_98 = lean_ctor_get(x_84, 11); -lean_inc(x_98); -x_99 = lean_ctor_get(x_84, 12); -lean_inc(x_99); -x_100 = lean_ctor_get(x_84, 13); -lean_inc(x_100); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - lean_ctor_release(x_84, 2); - lean_ctor_release(x_84, 3); - lean_ctor_release(x_84, 4); - lean_ctor_release(x_84, 5); - lean_ctor_release(x_84, 6); - lean_ctor_release(x_84, 7); - lean_ctor_release(x_84, 8); - lean_ctor_release(x_84, 9); - lean_ctor_release(x_84, 10); - lean_ctor_release(x_84, 11); - lean_ctor_release(x_84, 12); - lean_ctor_release(x_84, 13); - x_101 = x_84; -} else { - lean_dec_ref(x_84); - x_101 = lean_box(0); +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = lean_box(0); +lean_ctor_set(x_12, 0, x_17); +return x_12; } -x_102 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_97, x_1); -if (lean_is_scalar(x_101)) { - x_103 = lean_alloc_ctor(0, 14, 1); -} else { - x_103 = x_101; -} -lean_ctor_set(x_103, 0, x_86); -lean_ctor_set(x_103, 1, x_87); -lean_ctor_set(x_103, 2, x_88); -lean_ctor_set(x_103, 3, x_89); -lean_ctor_set(x_103, 4, x_90); -lean_ctor_set(x_103, 5, x_91); -lean_ctor_set(x_103, 6, x_93); -lean_ctor_set(x_103, 7, x_94); -lean_ctor_set(x_103, 8, x_95); -lean_ctor_set(x_103, 9, x_96); -lean_ctor_set(x_103, 10, x_102); -lean_ctor_set(x_103, 11, x_98); -lean_ctor_set(x_103, 12, x_99); -lean_ctor_set(x_103, 13, x_100); -lean_ctor_set_uint8(x_103, sizeof(void*)*14, x_92); -x_104 = lean_st_ref_set(x_3, x_103, x_85); -x_105 = lean_ctor_get(x_104, 1); -lean_inc(x_105); -lean_dec(x_104); -x_106 = lean_st_ref_get(x_3, x_105); -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); -lean_dec(x_106); -x_109 = lean_ctor_get(x_107, 4); -lean_inc(x_109); -lean_dec(x_107); -x_110 = lean_box(0); -x_111 = lean_box(0); -lean_inc(x_82); -x_112 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_2, x_82, x_109, x_110, x_82, x_82, x_111, lean_box(0), x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_108); -lean_dec(x_82); -if (lean_obj_tag(x_112) == 0) +else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_112, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_114 = x_112; -} else { - lean_dec_ref(x_112); - x_114 = lean_box(0); -} -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(0, 2, 0); -} else { - x_115 = x_114; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; } -lean_ctor_set(x_115, 0, x_111); -lean_ctor_set(x_115, 1, x_113); -return x_115; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_116 = lean_ctor_get(x_112, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_112, 1); -lean_inc(x_117); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_118 = x_112; -} else { - lean_dec_ref(x_112); - x_118 = lean_box(0); -} -if (lean_is_scalar(x_118)) { - x_119 = lean_alloc_ctor(1, 2, 0); -} else { - x_119 = x_118; -} -lean_ctor_set(x_119, 0, x_116); -lean_ctor_set(x_119, 1, x_117); -return x_119; -} -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3(x_1, x_2, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +return x_23; } } } @@ -11208,13 +15327,13 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_internalize___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_3); -return x_13; +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_internalize___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -11226,31 +15345,20 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_internalize___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_eraseAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_3); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_internalize___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_erase___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_2); -lean_dec(x_2); -return x_3; +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { uint8_t x_6; lean_object* x_7; -x_6 = l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); @@ -11258,50 +15366,33 @@ x_7 = lean_box(x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; uint8_t x_5; lean_object* x_6; x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__4(x_1, x_4, x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__2(x_1, x_4, x_3); lean_dec(x_3); x_6 = lean_box(x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__3(x_1, x_2); +x_3 = l_Lean_PersistentHashMap_contains___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__1(x_1, x_2); lean_dec(x_2); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -11314,7 +15405,7 @@ lean_dec(x_2); return x_12; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -11335,7 +15426,7 @@ lean_object* x_17 = _args[16]; _start: { lean_object* x_18; -x_18 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_18 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -11351,8 +15442,86 @@ lean_dec(x_1); return x_12; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_isMatcher___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_18 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___spec__2(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addMatchEqns___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +return x_13; +} +} lean_object* initialize_Init_Grind_Util(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_LitValues(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Match_MatcherInfo(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Match_MatchEqsExt(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -11363,9 +15532,18 @@ _G_initialized = true; res = initialize_Init_Grind_Util(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Grind_Lemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_LitValues(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Match_MatcherInfo(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Match_MatchEqsExt(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -11376,11 +15554,6 @@ l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec_ l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2 = _init_l_Lean_PersistentHashMap_findEntryAux___at_Lean_Meta_Grind_addCongrTable___spec__2___closed__2(); l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1(); lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_addCongrTable___spec__5___closed__1); -l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__1(); -l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__2); -l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9___closed__3); l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__1); l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__2(); @@ -11395,6 +15568,8 @@ l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6 = _init_l_Lean_Meta_Grin lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__6); l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__7); +l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8 = _init_l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___lambda__2___closed__8); l_Lean_Meta_Grind_addCongrTable___closed__1 = _init_l_Lean_Meta_Grind_addCongrTable___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__1); l_Lean_Meta_Grind_addCongrTable___closed__2 = _init_l_Lean_Meta_Grind_addCongrTable___closed__2(); @@ -11413,6 +15588,74 @@ l_Lean_Meta_Grind_addCongrTable___closed__8 = _init_l_Lean_Meta_Grind_addCongrTa lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__8); l_Lean_Meta_Grind_addCongrTable___closed__9 = _init_l_Lean_Meta_Grind_addCongrTable___closed__9(); lean_mark_persistent(l_Lean_Meta_Grind_addCongrTable___closed__9); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_addSplitCandidate___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__4); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__5); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__6); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__7); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__8); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__9); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__10); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__11); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes___closed__12); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_forbiddenSplitTypes); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___lambda__4___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__4); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__5); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__6); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__7); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_pushCastHEqs___closed__8); l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__1); l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_internalizePattern___closed__2(); @@ -11437,48 +15680,44 @@ l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1 = lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__1); l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2(); lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__2); -l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__3); -l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4 = _init_l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4(); -lean_mark_persistent(l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_internalize___spec__3___closed__4); -l_Lean_Meta_Grind_internalize___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__1); -l_Lean_Meta_Grind_internalize___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__2); -l_Lean_Meta_Grind_internalize___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__3); -l_Lean_Meta_Grind_internalize___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__4); -l_Lean_Meta_Grind_internalize___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__5); -l_Lean_Meta_Grind_internalize___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_internalize___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__2___closed__6); l_Lean_Meta_Grind_internalize___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__1); l_Lean_Meta_Grind_internalize___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___lambda__1___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__4); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__5); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__6); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__7); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__8); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__9); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__8___closed__10); +l_Lean_Meta_Grind_internalize___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__3); +l_Lean_Meta_Grind_internalize___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__4); +l_Lean_Meta_Grind_internalize___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__5); +l_Lean_Meta_Grind_internalize___lambda__3___closed__6 = _init_l_Lean_Meta_Grind_internalize___lambda__3___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__3___closed__6); +l_Lean_Meta_Grind_internalize___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_internalize___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__4___closed__1); +l_Lean_Meta_Grind_internalize___lambda__4___closed__2 = _init_l_Lean_Meta_Grind_internalize___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_internalize___lambda__4___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___lambda__1___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__5); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheoremPatterns___spec__5___closed__6); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__1); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__2); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__3); +l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_activateTheorem___closed__4); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c index 1ccf0922ead3..d6e02cc552d3 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Intro.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Intro -// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.Injection Lean.Meta.Tactic.Grind.Core Lean.Meta.Tactic.Grind.Combinators #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,42 +15,47 @@ extern "C" { #endif lean_object* l_Lean_Expr_bindingName_x21(lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_injection_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedIntroResult; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshExprMVarAt(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLet(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isLetFun(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertNext_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Meta_mkEqMP(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_add(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_addNewEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_assertAll___closed__1; +lean_object* l_Lean_LocalDecl_value(lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_intros___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -59,57 +64,59 @@ lean_object* lean_array_to_list(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_assertAt___closed__1; lean_object* l_Lean_Meta_Grind_isGrindCasesTarget(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9___boxed(lean_object**); lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Queue_dequeue_x3f___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); uint8_t l_Lean_Expr_isArrow(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_assertAt___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); lean_object* l_Lean_Meta_Grind_addHypothesis(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2; lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_intros_go___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed(lean_object**); lean_object* l_Lean_LocalDecl_type(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__4___boxed(lean_object**); lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_byContra_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); lean_object* l_Lean_Expr_fvar___override(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___boxed(lean_object**); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_cases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4; lean_object* l_Lean_Expr_bindingInfo_x21(lean_object*); uint8_t l_Lean_Expr_isForall(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -118,9 +125,9 @@ lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Meta_Grind_instInhabitedIntroResult() { _start: { @@ -559,1390 +566,1905 @@ x_10 = l_Lean_FVarId_getDecl(x_1, x_5, x_6, x_7, x_8, x_9); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_26; lean_object* x_27; -x_26 = l_Lean_LocalDecl_type(x_17); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_26); -x_27 = l_Lean_Meta_isProp(x_26, x_21, x_22, x_23, x_24, x_25); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; uint8_t x_29; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_unbox(x_28); -lean_dec(x_28); -if (x_29 == 0) -{ -uint8_t x_30; -lean_dec(x_26); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -x_30 = !lean_is_exclusive(x_27); -if (x_30 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_27, 0); -lean_dec(x_31); -x_32 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_32, 0, x_1); -lean_ctor_set(x_32, 1, x_2); -lean_ctor_set(x_32, 2, x_3); -lean_ctor_set(x_32, 3, x_4); -lean_ctor_set(x_32, 4, x_5); -lean_ctor_set(x_32, 5, x_6); -lean_ctor_set(x_32, 6, x_8); -lean_ctor_set(x_32, 7, x_9); -lean_ctor_set(x_32, 8, x_10); -lean_ctor_set(x_32, 9, x_11); -lean_ctor_set(x_32, 10, x_12); -lean_ctor_set(x_32, 11, x_13); -lean_ctor_set(x_32, 12, x_14); -lean_ctor_set(x_32, 13, x_15); -lean_ctor_set_uint8(x_32, sizeof(void*)*14, x_7); -x_33 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_33, 0, x_16); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_27, 0, x_33); -return x_27; +return x_10; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); -lean_dec(x_27); -x_35 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_35, 0, x_1); -lean_ctor_set(x_35, 1, x_2); -lean_ctor_set(x_35, 2, x_3); -lean_ctor_set(x_35, 3, x_4); -lean_ctor_set(x_35, 4, x_5); -lean_ctor_set(x_35, 5, x_6); -lean_ctor_set(x_35, 6, x_8); -lean_ctor_set(x_35, 7, x_9); -lean_ctor_set(x_35, 8, x_10); -lean_ctor_set(x_35, 9, x_11); -lean_ctor_set(x_35, 10, x_12); -lean_ctor_set(x_35, 11, x_13); -lean_ctor_set(x_35, 12, x_14); -lean_ctor_set(x_35, 13, x_15); -lean_ctor_set_uint8(x_35, sizeof(void*)*14, x_7); -x_36 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_36, 0, x_16); -lean_ctor_set(x_36, 1, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_34); -return x_37; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_38 = lean_ctor_get(x_27, 1); -lean_inc(x_38); -lean_dec(x_27); -x_39 = l_Lean_LocalDecl_userName(x_17); -x_40 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_39, x_23, x_24, x_38); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_Expr_fvar___override(x_16); -x_44 = l_Lean_MVarId_assert(x_1, x_41, x_26, x_43, x_21, x_22, x_23, x_24, x_42); -if (lean_obj_tag(x_44) == 0) -{ -uint8_t x_45; -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_44, 0); -x_47 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_2); -lean_ctor_set(x_47, 2, x_3); -lean_ctor_set(x_47, 3, x_4); -lean_ctor_set(x_47, 4, x_5); -lean_ctor_set(x_47, 5, x_6); -lean_ctor_set(x_47, 6, x_8); -lean_ctor_set(x_47, 7, x_9); -lean_ctor_set(x_47, 8, x_10); -lean_ctor_set(x_47, 9, x_11); -lean_ctor_set(x_47, 10, x_12); -lean_ctor_set(x_47, 11, x_13); -lean_ctor_set(x_47, 12, x_14); -lean_ctor_set(x_47, 13, x_15); -lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_7); -x_48 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_44, 0, x_48); -return x_44; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_49 = lean_ctor_get(x_44, 0); -x_50 = lean_ctor_get(x_44, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_44); -x_51 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_2); -lean_ctor_set(x_51, 2, x_3); -lean_ctor_set(x_51, 3, x_4); -lean_ctor_set(x_51, 4, x_5); -lean_ctor_set(x_51, 5, x_6); -lean_ctor_set(x_51, 6, x_8); -lean_ctor_set(x_51, 7, x_9); -lean_ctor_set(x_51, 8, x_10); -lean_ctor_set(x_51, 9, x_11); -lean_ctor_set(x_51, 10, x_12); -lean_ctor_set(x_51, 11, x_13); -lean_ctor_set(x_51, 12, x_14); -lean_ctor_set(x_51, 13, x_15); -lean_ctor_set_uint8(x_51, sizeof(void*)*14, x_7); -x_52 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_52, 0, x_51); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_50); -return x_53; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_54; -lean_dec(x_15); +lean_object* x_14; +lean_inc(x_5); +x_14 = l_Lean_Meta_Grind_addNewEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); lean_dec(x_14); -lean_dec(x_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); +x_16 = lean_st_ref_get(x_5, x_15); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_16, 1); +x_19 = lean_st_ref_get(x_5, x_18); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_54 = !lean_is_exclusive(x_44); -if (x_54 == 0) +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -return x_44; +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_ctor_set(x_16, 1, x_21); +lean_ctor_set(x_19, 0, x_16); +return x_19; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_44, 0); -x_56 = lean_ctor_get(x_44, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_44); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_19, 0); +x_23 = lean_ctor_get(x_19, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_19); +lean_ctor_set(x_16, 1, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_16); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} } +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_25 = lean_ctor_get(x_16, 0); +x_26 = lean_ctor_get(x_16, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_16); +x_27 = lean_st_ref_get(x_5, x_26); +lean_dec(x_5); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_30 = x_27; +} else { + lean_dec_ref(x_27); + x_30 = lean_box(0); +} +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_25); +lean_ctor_set(x_31, 1, x_28); +if (lean_is_scalar(x_30)) { + x_32 = lean_alloc_ctor(0, 2, 0); +} else { + x_32 = x_30; } +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_29); +return x_32; } } else { -uint8_t x_58; -lean_dec(x_26); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); +uint8_t x_33; lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_58 = !lean_is_exclusive(x_27); -if (x_58 == 0) +x_33 = !lean_is_exclusive(x_14); +if (x_33 == 0) { -return x_27; +return x_14; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_27, 0); -x_60 = lean_ctor_get(x_27, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_27); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; -} -} +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_14, 0); +x_35 = lean_ctor_get(x_14, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_14); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Meta_isProp(x_1, x_5, x_6, x_7, x_8, x_9); -return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; -x_11 = 0; -x_12 = 1; -x_13 = 1; -x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_11, x_12, x_11, x_13, x_6, x_7, x_8, x_9, x_10); -return x_14; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1() { -_start: +uint8_t x_10; +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); +lean_object* x_11; +x_11 = lean_ctor_get(x_1, 1); +lean_dec(x_11); +lean_ctor_set(x_1, 1, x_9); return x_1; } -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Grind", 5, 5); -return x_1; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("intro_with_eq", 13, 13); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed), 9, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2; -x_3 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, uint8_t x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31, lean_object* x_32, lean_object* x_33) { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_34; lean_object* x_35; +x_34 = l_Lean_LocalDecl_type(x_25); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_34); +x_35 = l_Lean_Meta_isProp(x_34, x_29, x_30, x_31, x_32, x_33); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_32; uint8_t x_33; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -x_32 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) +lean_object* x_36; uint8_t x_37; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_unbox(x_36); +lean_dec(x_36); +if (x_37 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_32, 0); +uint8_t x_38; lean_dec(x_34); -x_35 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_35, 0, x_3); -lean_ctor_set(x_35, 1, x_4); -lean_ctor_set(x_35, 2, x_5); -lean_ctor_set(x_35, 3, x_6); -lean_ctor_set(x_35, 4, x_7); -lean_ctor_set(x_35, 5, x_8); -lean_ctor_set(x_35, 6, x_10); -lean_ctor_set(x_35, 7, x_11); -lean_ctor_set(x_35, 8, x_12); -lean_ctor_set(x_35, 9, x_13); -lean_ctor_set(x_35, 10, x_14); -lean_ctor_set(x_35, 11, x_15); -lean_ctor_set(x_35, 12, x_16); -lean_ctor_set(x_35, 13, x_17); -lean_ctor_set_uint8(x_35, sizeof(void*)*14, x_9); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_18); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set(x_32, 0, x_36); -return x_32; +x_38 = !lean_is_exclusive(x_35); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_35, 1); +x_40 = lean_ctor_get(x_35, 0); +lean_dec(x_40); +lean_inc(x_1); +x_41 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_41, 0, x_1); +lean_ctor_set(x_41, 1, x_2); +lean_ctor_set(x_41, 2, x_3); +lean_ctor_set(x_41, 3, x_4); +lean_ctor_set(x_41, 4, x_5); +lean_ctor_set(x_41, 5, x_6); +lean_ctor_set(x_41, 6, x_8); +lean_ctor_set(x_41, 7, x_9); +lean_ctor_set(x_41, 8, x_10); +lean_ctor_set(x_41, 9, x_11); +lean_ctor_set(x_41, 10, x_12); +lean_ctor_set(x_41, 11, x_13); +lean_ctor_set(x_41, 12, x_14); +lean_ctor_set(x_41, 13, x_15); +lean_ctor_set(x_41, 14, x_16); +lean_ctor_set(x_41, 15, x_17); +lean_ctor_set(x_41, 16, x_18); +lean_ctor_set(x_41, 17, x_19); +lean_ctor_set(x_41, 18, x_20); +lean_ctor_set_uint8(x_41, sizeof(void*)*19, x_7); +if (x_23 == 0) +{ +uint8_t x_112; +x_112 = l_Lean_Expr_isLetFun(x_24); +if (x_112 == 0) +{ +lean_object* x_113; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_1); +x_113 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_113, 0, x_21); +lean_ctor_set(x_113, 1, x_41); +lean_ctor_set(x_35, 0, x_113); +return x_35; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = lean_ctor_get(x_32, 1); -lean_inc(x_37); -lean_dec(x_32); -x_38 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_38, 0, x_3); -lean_ctor_set(x_38, 1, x_4); -lean_ctor_set(x_38, 2, x_5); -lean_ctor_set(x_38, 3, x_6); -lean_ctor_set(x_38, 4, x_7); -lean_ctor_set(x_38, 5, x_8); -lean_ctor_set(x_38, 6, x_10); -lean_ctor_set(x_38, 7, x_11); -lean_ctor_set(x_38, 8, x_12); -lean_ctor_set(x_38, 9, x_13); -lean_ctor_set(x_38, 10, x_14); -lean_ctor_set(x_38, 11, x_15); -lean_ctor_set(x_38, 12, x_16); -lean_ctor_set(x_38, 13, x_17); -lean_ctor_set_uint8(x_38, sizeof(void*)*14, x_9); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_18); -lean_ctor_set(x_39, 1, x_38); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_37); -return x_40; +lean_object* x_114; +lean_free_object(x_35); +x_114 = lean_box(0); +x_42 = x_114; +goto block_111; } } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_41 = lean_ctor_get(x_1, 0); -x_42 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4; -lean_inc(x_19); -x_43 = l_Lean_Expr_const___override(x_42, x_19); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_23); -lean_ctor_set(x_44, 1, x_19); -lean_inc(x_41); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_41); -lean_ctor_set(x_45, 1, x_44); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_20); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_21); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_22); -lean_ctor_set(x_48, 1, x_47); -x_49 = lean_array_mk(x_48); -x_50 = l_Lean_mkAppN(x_43, x_49); -lean_dec(x_49); -x_51 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_50, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); +lean_object* x_115; +lean_free_object(x_35); +x_115 = lean_box(0); +x_42 = x_115; +goto block_111; +} +block_111: +{ +lean_object* x_43; +lean_dec(x_42); +lean_inc(x_29); +lean_inc(x_21); +x_43 = l_Lean_FVarId_getDecl(x_21, x_29, x_30, x_31, x_32, x_39); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_LocalDecl_value(x_44); +lean_dec(x_44); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_27); +x_47 = l_Lean_Meta_Grind_simp(x_46, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_45); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t 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); +lean_inc(x_21); +x_50 = l_Lean_Expr_fvar___override(x_21); +x_51 = l_Lean_Meta_Grind_shareCommon(x_50, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_49); x_52 = !lean_is_exclusive(x_51); if (x_52 == 0) { lean_object* x_53; lean_object* x_54; lean_object* x_55; x_53 = lean_ctor_get(x_51, 0); -lean_dec(x_53); -x_54 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_54, 0, x_3); -lean_ctor_set(x_54, 1, x_4); -lean_ctor_set(x_54, 2, x_5); -lean_ctor_set(x_54, 3, x_6); -lean_ctor_set(x_54, 4, x_7); -lean_ctor_set(x_54, 5, x_8); -lean_ctor_set(x_54, 6, x_10); -lean_ctor_set(x_54, 7, x_11); -lean_ctor_set(x_54, 8, x_12); -lean_ctor_set(x_54, 9, x_13); -lean_ctor_set(x_54, 10, x_14); -lean_ctor_set(x_54, 11, x_15); -lean_ctor_set(x_54, 12, x_16); -lean_ctor_set(x_54, 13, x_17); -lean_ctor_set_uint8(x_54, sizeof(void*)*14, x_9); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_18); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_51, 0, x_55); -return x_51; -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_51, 1); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_48); +x_55 = l_Lean_Meta_Simp_Result_getProof(x_48, x_29, x_30, x_31, x_32, x_54); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_56 = lean_ctor_get(x_55, 0); lean_inc(x_56); -lean_dec(x_51); -x_57 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_57, 0, x_3); -lean_ctor_set(x_57, 1, x_4); -lean_ctor_set(x_57, 2, x_5); -lean_ctor_set(x_57, 3, x_6); -lean_ctor_set(x_57, 4, x_7); -lean_ctor_set(x_57, 5, x_8); -lean_ctor_set(x_57, 6, x_10); -lean_ctor_set(x_57, 7, x_11); -lean_ctor_set(x_57, 8, x_12); -lean_ctor_set(x_57, 9, x_13); -lean_ctor_set(x_57, 10, x_14); -lean_ctor_set(x_57, 11, x_15); -lean_ctor_set(x_57, 12, x_16); -lean_ctor_set(x_57, 13, x_17); -lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_9); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_18); -lean_ctor_set(x_58, 1, x_57); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_56); -return x_59; -} +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_ctor_get(x_48, 0); +lean_inc(x_58); +lean_dec(x_48); +x_59 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_59, 0, x_41); +x_60 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3), 13, 4); +lean_closure_set(x_60, 0, x_53); +lean_closure_set(x_60, 1, x_58); +lean_closure_set(x_60, 2, x_56); +lean_closure_set(x_60, 3, x_22); +x_61 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_61, 0, x_59); +lean_closure_set(x_61, 1, x_60); +x_62 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_63 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_63, 0, x_61); +lean_closure_set(x_63, 1, x_62); +x_64 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_63, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_57); +if (lean_obj_tag(x_64) == 0) +{ +uint8_t x_65; +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) +{ +lean_object* x_66; +x_66 = lean_ctor_get(x_64, 0); +lean_ctor_set_tag(x_51, 3); +lean_ctor_set(x_51, 1, x_66); +lean_ctor_set(x_51, 0, x_21); +lean_ctor_set(x_64, 0, x_51); +return x_64; } +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_64, 0); +x_68 = lean_ctor_get(x_64, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_64); +lean_ctor_set_tag(x_51, 3); +lean_ctor_set(x_51, 1, x_67); +lean_ctor_set(x_51, 0, x_21); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_51); +lean_ctor_set(x_69, 1, x_68); +return x_69; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, uint8_t x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26) { -_start: -{ -if (x_18 == 0) +else { -uint8_t x_27; lean_object* x_28; +uint8_t x_70; +lean_free_object(x_51); lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_17); -x_27 = 1; -x_28 = l_Lean_Meta_intro1Core(x_1, x_27, x_22, x_23, x_24, x_25, x_26); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_28, 0); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_2); -lean_ctor_set(x_33, 2, x_3); -lean_ctor_set(x_33, 3, x_4); -lean_ctor_set(x_33, 4, x_5); -lean_ctor_set(x_33, 5, x_6); -lean_ctor_set(x_33, 6, x_8); -lean_ctor_set(x_33, 7, x_9); -lean_ctor_set(x_33, 8, x_10); -lean_ctor_set(x_33, 9, x_11); -lean_ctor_set(x_33, 10, x_12); -lean_ctor_set(x_33, 11, x_13); -lean_ctor_set(x_33, 12, x_14); -lean_ctor_set(x_33, 13, x_15); -lean_ctor_set_uint8(x_33, sizeof(void*)*14, x_7); -lean_ctor_set_tag(x_30, 3); -lean_ctor_set(x_30, 1, x_33); -return x_28; -} -else +x_70 = !lean_is_exclusive(x_64); +if (x_70 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_30); -x_36 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_2); -lean_ctor_set(x_36, 2, x_3); -lean_ctor_set(x_36, 3, x_4); -lean_ctor_set(x_36, 4, x_5); -lean_ctor_set(x_36, 5, x_6); -lean_ctor_set(x_36, 6, x_8); -lean_ctor_set(x_36, 7, x_9); -lean_ctor_set(x_36, 8, x_10); -lean_ctor_set(x_36, 9, x_11); -lean_ctor_set(x_36, 10, x_12); -lean_ctor_set(x_36, 11, x_13); -lean_ctor_set(x_36, 12, x_14); -lean_ctor_set(x_36, 13, x_15); -lean_ctor_set_uint8(x_36, sizeof(void*)*14, x_7); -x_37 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; -} +return x_64; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_38 = lean_ctor_get(x_28, 0); -x_39 = lean_ctor_get(x_28, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_28); -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; -} else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -x_43 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_2); -lean_ctor_set(x_43, 2, x_3); -lean_ctor_set(x_43, 3, x_4); -lean_ctor_set(x_43, 4, x_5); -lean_ctor_set(x_43, 5, x_6); -lean_ctor_set(x_43, 6, x_8); -lean_ctor_set(x_43, 7, x_9); -lean_ctor_set(x_43, 8, x_10); -lean_ctor_set(x_43, 9, x_11); -lean_ctor_set(x_43, 10, x_12); -lean_ctor_set(x_43, 11, x_13); -lean_ctor_set(x_43, 12, x_14); -lean_ctor_set(x_43, 13, x_15); -lean_ctor_set_uint8(x_43, sizeof(void*)*14, x_7); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(3, 2, 0); -} else { - x_44 = x_42; - lean_ctor_set_tag(x_44, 3); +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_64, 0); +x_72 = lean_ctor_get(x_64, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_64); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +return x_73; } -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_39); -return x_45; } } else { -uint8_t x_46; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_46 = !lean_is_exclusive(x_28); -if (x_46 == 0) +uint8_t x_74; +lean_free_object(x_51); +lean_dec(x_53); +lean_dec(x_48); +lean_dec(x_41); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_1); +x_74 = !lean_is_exclusive(x_55); +if (x_74 == 0) { -return x_28; +return x_55; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 0); -x_48 = lean_ctor_get(x_28, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_28); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_55, 0); +x_76 = lean_ctor_get(x_55, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_55); +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +return x_77; } } } else { -lean_object* x_50; -lean_inc(x_1); -x_50 = l_Lean_MVarId_getTag(x_1, x_22, x_23, x_24, x_25, x_26); -if (lean_obj_tag(x_50) == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); -x_53 = l_Lean_Expr_bindingBody_x21(x_16); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_20); -lean_inc(x_17); -x_54 = l_Lean_Meta_Grind_simp(x_17, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_52); -if (lean_obj_tag(x_54) == 0) +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_51, 0); +x_79 = lean_ctor_get(x_51, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_51); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_48); +x_80 = l_Lean_Meta_Simp_Result_getProof(x_48, x_29, x_30, x_31, x_32, x_79); +if (lean_obj_tag(x_80) == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -x_57 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1(x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_56); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_60 = lean_ctor_get(x_22, 2); -lean_inc(x_60); -x_61 = l_Lean_Expr_bindingName_x21(x_16); -x_62 = lean_ctor_get(x_55, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_55, 1); -lean_inc(x_63); -lean_dec(x_55); -x_64 = l_Lean_Expr_bindingInfo_x21(x_16); -x_65 = lean_unbox(x_64); -lean_dec(x_64); -x_66 = 0; -lean_inc(x_62); -lean_inc(x_58); -x_67 = l_Lean_LocalContext_mkLocalDecl(x_60, x_58, x_61, x_62, x_65, x_66); -x_68 = l_Lean_Meta_getLocalInstances(x_22, x_23, x_24, x_25, x_59); -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = 2; -x_72 = lean_unsigned_to_nat(0u); -lean_inc(x_53); -x_73 = l_Lean_Meta_mkFreshExprMVarAt(x_67, x_69, x_53, x_71, x_51, x_72, x_22, x_23, x_24, x_25, x_70); -x_74 = !lean_is_exclusive(x_73); -if (x_74 == 0) +lean_object* x_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_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +lean_dec(x_80); +x_83 = lean_ctor_get(x_48, 0); +lean_inc(x_83); +lean_dec(x_48); +x_84 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_84, 0, x_41); +x_85 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3), 13, 4); +lean_closure_set(x_85, 0, x_78); +lean_closure_set(x_85, 1, x_83); +lean_closure_set(x_85, 2, x_81); +lean_closure_set(x_85, 3, x_22); +x_86 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_86, 0, x_84); +lean_closure_set(x_86, 1, x_85); +x_87 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_88 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_88, 0, x_86); +lean_closure_set(x_88, 1, x_87); +x_89 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_88, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_82); +if (lean_obj_tag(x_89) == 0) { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_75 = lean_ctor_get(x_73, 0); -x_76 = lean_ctor_get(x_73, 1); -x_77 = l_Lean_Expr_mvarId_x21(x_75); -lean_inc(x_58); -x_78 = l_Lean_Expr_fvar___override(x_58); -x_79 = lean_box(0); -lean_ctor_set_tag(x_73, 1); -lean_ctor_set(x_73, 1, x_79); -lean_ctor_set(x_73, 0, x_78); -x_80 = lean_array_mk(x_73); -x_81 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed), 10, 2); -lean_closure_set(x_81, 0, x_80); -lean_closure_set(x_81, 1, x_75); -x_82 = lean_box(x_7); -lean_inc(x_77); -x_83 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 31, 22); -lean_closure_set(x_83, 0, x_63); -lean_closure_set(x_83, 1, x_1); -lean_closure_set(x_83, 2, x_77); -lean_closure_set(x_83, 3, x_2); -lean_closure_set(x_83, 4, x_3); -lean_closure_set(x_83, 5, x_4); -lean_closure_set(x_83, 6, x_5); -lean_closure_set(x_83, 7, x_6); -lean_closure_set(x_83, 8, x_82); -lean_closure_set(x_83, 9, x_8); -lean_closure_set(x_83, 10, x_9); -lean_closure_set(x_83, 11, x_10); -lean_closure_set(x_83, 12, x_11); -lean_closure_set(x_83, 13, x_12); -lean_closure_set(x_83, 14, x_13); -lean_closure_set(x_83, 15, x_14); -lean_closure_set(x_83, 16, x_15); -lean_closure_set(x_83, 17, x_58); -lean_closure_set(x_83, 18, x_79); -lean_closure_set(x_83, 19, x_53); -lean_closure_set(x_83, 20, x_62); -lean_closure_set(x_83, 21, x_17); -x_84 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_84, 0, x_81); -lean_closure_set(x_84, 1, x_83); -x_85 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_77, x_84, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_76); -return x_85; -} -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_86 = lean_ctor_get(x_73, 0); -x_87 = lean_ctor_get(x_73, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_73); -x_88 = l_Lean_Expr_mvarId_x21(x_86); -lean_inc(x_58); -x_89 = l_Lean_Expr_fvar___override(x_58); -x_90 = lean_box(0); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_92 = lean_array_mk(x_91); -x_93 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed), 10, 2); -lean_closure_set(x_93, 0, x_92); -lean_closure_set(x_93, 1, x_86); -x_94 = lean_box(x_7); -lean_inc(x_88); -x_95 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 31, 22); -lean_closure_set(x_95, 0, x_63); -lean_closure_set(x_95, 1, x_1); -lean_closure_set(x_95, 2, x_88); -lean_closure_set(x_95, 3, x_2); -lean_closure_set(x_95, 4, x_3); -lean_closure_set(x_95, 5, x_4); -lean_closure_set(x_95, 6, x_5); -lean_closure_set(x_95, 7, x_6); -lean_closure_set(x_95, 8, x_94); -lean_closure_set(x_95, 9, x_8); -lean_closure_set(x_95, 10, x_9); -lean_closure_set(x_95, 11, x_10); -lean_closure_set(x_95, 12, x_11); -lean_closure_set(x_95, 13, x_12); -lean_closure_set(x_95, 14, x_13); -lean_closure_set(x_95, 15, x_14); -lean_closure_set(x_95, 16, x_15); -lean_closure_set(x_95, 17, x_58); -lean_closure_set(x_95, 18, x_90); -lean_closure_set(x_95, 19, x_53); -lean_closure_set(x_95, 20, x_62); -lean_closure_set(x_95, 21, x_17); -x_96 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_96, 0, x_93); -lean_closure_set(x_96, 1, x_95); -x_97 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_88, x_96, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_87); -return x_97; +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_92 = x_89; +} else { + lean_dec_ref(x_89); + x_92 = lean_box(0); +} +x_93 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_93, 0, x_21); +lean_ctor_set(x_93, 1, x_90); +if (lean_is_scalar(x_92)) { + x_94 = lean_alloc_ctor(0, 2, 0); +} else { + x_94 = x_92; } +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_91); +return x_94; } else { -uint8_t x_98; -lean_dec(x_53); -lean_dec(x_51); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_17); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_98 = !lean_is_exclusive(x_54); -if (x_98 == 0) -{ -return x_54; +x_95 = lean_ctor_get(x_89, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_89, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_97 = x_89; +} else { + lean_dec_ref(x_89); + x_97 = lean_box(0); +} +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(1, 2, 0); +} else { + x_98 = x_97; +} +lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 1, x_96); +return x_98; +} } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_54, 0); -x_100 = lean_ctor_get(x_54, 1); -lean_inc(x_100); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_78); +lean_dec(x_48); +lean_dec(x_41); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_1); +x_99 = lean_ctor_get(x_80, 0); lean_inc(x_99); -lean_dec(x_54); -x_101 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_101, 0, x_99); -lean_ctor_set(x_101, 1, x_100); -return x_101; +x_100 = lean_ctor_get(x_80, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_101 = x_80; +} else { + lean_dec_ref(x_80); + x_101 = lean_box(0); +} +if (lean_is_scalar(x_101)) { + x_102 = lean_alloc_ctor(1, 2, 0); +} else { + x_102 = x_101; +} +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_100); +return x_102; } } } else { -uint8_t x_102; -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); +uint8_t x_103; +lean_dec(x_41); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); lean_dec(x_22); lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_17); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_102 = !lean_is_exclusive(x_50); -if (x_102 == 0) +x_103 = !lean_is_exclusive(x_47); +if (x_103 == 0) { -return x_50; +return x_47; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_103 = lean_ctor_get(x_50, 0); -x_104 = lean_ctor_get(x_50, 1); +lean_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_inc(x_103); -lean_dec(x_50); -x_105 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -return x_105; -} -} +lean_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; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +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; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 2); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 3); -lean_inc(x_13); -x_14 = lean_ctor_get(x_1, 4); -lean_inc(x_14); -x_15 = lean_ctor_get(x_1, 5); -lean_inc(x_15); -x_16 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_17 = lean_ctor_get(x_1, 6); -lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 7); -lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 8); -lean_inc(x_19); -x_20 = lean_ctor_get(x_1, 9); -lean_inc(x_20); -x_21 = lean_ctor_get(x_1, 10); -lean_inc(x_21); -x_22 = lean_ctor_get(x_1, 11); -lean_inc(x_22); -x_23 = lean_ctor_get(x_1, 12); -lean_inc(x_23); -x_24 = lean_ctor_get(x_1, 13); -lean_inc(x_24); +uint8_t x_107; +lean_dec(x_41); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_21); lean_dec(x_1); -lean_inc(x_10); -x_25 = l_Lean_MVarId_getType(x_10, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_25) == 0) -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +x_107 = !lean_is_exclusive(x_43); +if (x_107 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_46; -x_27 = lean_ctor_get(x_25, 0); -x_28 = lean_ctor_get(x_25, 1); -x_46 = l_Lean_Expr_isArrow(x_27); -if (x_46 == 0) +return x_43; +} +else { -uint8_t x_47; -x_47 = l_Lean_Expr_isLet(x_27); -if (x_47 == 0) +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_43, 0); +x_109 = lean_ctor_get(x_43, 1); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_43); +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 { -uint8_t x_48; -x_48 = l_Lean_Expr_isForall(x_27); +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_35, 1); +lean_inc(x_116); +lean_dec(x_35); +lean_inc(x_1); +x_117 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_117, 0, x_1); +lean_ctor_set(x_117, 1, x_2); +lean_ctor_set(x_117, 2, x_3); +lean_ctor_set(x_117, 3, x_4); +lean_ctor_set(x_117, 4, x_5); +lean_ctor_set(x_117, 5, x_6); +lean_ctor_set(x_117, 6, x_8); +lean_ctor_set(x_117, 7, x_9); +lean_ctor_set(x_117, 8, x_10); +lean_ctor_set(x_117, 9, x_11); +lean_ctor_set(x_117, 10, x_12); +lean_ctor_set(x_117, 11, x_13); +lean_ctor_set(x_117, 12, x_14); +lean_ctor_set(x_117, 13, x_15); +lean_ctor_set(x_117, 14, x_16); +lean_ctor_set(x_117, 15, x_17); +lean_ctor_set(x_117, 16, x_18); +lean_ctor_set(x_117, 17, x_19); +lean_ctor_set(x_117, 18, x_20); +lean_ctor_set_uint8(x_117, sizeof(void*)*19, x_7); +if (x_23 == 0) +{ +uint8_t x_163; +x_163 = l_Lean_Expr_isLetFun(x_24); +if (x_163 == 0) +{ +lean_object* x_164; lean_object* x_165; +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); lean_dec(x_27); -if (x_48 == 0) -{ -lean_object* x_49; -lean_dec(x_24); -lean_dec(x_23); +lean_dec(x_26); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_49 = lean_box(0); -lean_ctor_set(x_25, 0, x_49); -return x_25; +lean_dec(x_1); +x_164 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_164, 0, x_21); +lean_ctor_set(x_164, 1, x_117); +x_165 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_165, 0, x_164); +lean_ctor_set(x_165, 1, x_116); +return x_165; } else { -lean_object* x_50; -lean_free_object(x_25); -x_50 = lean_box(0); -x_29 = x_50; -goto block_45; +lean_object* x_166; +x_166 = lean_box(0); +x_118 = x_166; +goto block_162; } } else { -lean_object* x_51; -lean_free_object(x_25); -lean_dec(x_27); -x_51 = lean_box(0); -x_29 = x_51; -goto block_45; +lean_object* x_167; +x_167 = lean_box(0); +x_118 = x_167; +goto block_162; +} +block_162: +{ +lean_object* x_119; +lean_dec(x_118); +lean_inc(x_29); +lean_inc(x_21); +x_119 = l_Lean_FVarId_getDecl(x_21, x_29, x_30, x_31, x_32, x_116); +if (lean_obj_tag(x_119) == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_119, 1); +lean_inc(x_121); +lean_dec(x_119); +x_122 = l_Lean_LocalDecl_value(x_120); +lean_dec(x_120); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_27); +x_123 = l_Lean_Meta_Grind_simp(x_122, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_121); +if (lean_obj_tag(x_123) == 0) +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +lean_dec(x_123); +lean_inc(x_21); +x_126 = l_Lean_Expr_fvar___override(x_21); +x_127 = l_Lean_Meta_Grind_shareCommon(x_126, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_125); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_130 = x_127; +} else { + lean_dec_ref(x_127); + x_130 = lean_box(0); +} +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_124); +x_131 = l_Lean_Meta_Simp_Result_getProof(x_124, x_29, x_30, x_31, x_32, x_129); +if (lean_obj_tag(x_131) == 0) +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_131, 1); +lean_inc(x_133); +lean_dec(x_131); +x_134 = lean_ctor_get(x_124, 0); +lean_inc(x_134); +lean_dec(x_124); +x_135 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_135, 0, x_117); +x_136 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3), 13, 4); +lean_closure_set(x_136, 0, x_128); +lean_closure_set(x_136, 1, x_134); +lean_closure_set(x_136, 2, x_132); +lean_closure_set(x_136, 3, x_22); +x_137 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_137, 0, x_135); +lean_closure_set(x_137, 1, x_136); +x_138 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_139 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_139, 0, x_137); +lean_closure_set(x_139, 1, x_138); +x_140 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_139, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_133); +if (lean_obj_tag(x_140) == 0) +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_141 = lean_ctor_get(x_140, 0); +lean_inc(x_141); +x_142 = lean_ctor_get(x_140, 1); +lean_inc(x_142); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_143 = x_140; +} else { + lean_dec_ref(x_140); + x_143 = lean_box(0); } +if (lean_is_scalar(x_130)) { + x_144 = lean_alloc_ctor(3, 2, 0); +} else { + x_144 = x_130; + lean_ctor_set_tag(x_144, 3); +} +lean_ctor_set(x_144, 0, x_21); +lean_ctor_set(x_144, 1, x_141); +if (lean_is_scalar(x_143)) { + x_145 = lean_alloc_ctor(0, 2, 0); +} else { + x_145 = x_143; +} +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_142); +return x_145; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_free_object(x_25); -x_52 = l_Lean_Expr_bindingDomain_x21(x_27); -lean_inc(x_52); -x_53 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed), 9, 1); -lean_closure_set(x_53, 0, x_52); -x_54 = lean_box(x_16); -lean_inc(x_10); -x_55 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 26, 17); -lean_closure_set(x_55, 0, x_10); -lean_closure_set(x_55, 1, x_11); -lean_closure_set(x_55, 2, x_12); -lean_closure_set(x_55, 3, x_13); -lean_closure_set(x_55, 4, x_14); -lean_closure_set(x_55, 5, x_15); -lean_closure_set(x_55, 6, x_54); -lean_closure_set(x_55, 7, x_17); -lean_closure_set(x_55, 8, x_18); -lean_closure_set(x_55, 9, x_19); -lean_closure_set(x_55, 10, x_20); -lean_closure_set(x_55, 11, x_21); -lean_closure_set(x_55, 12, x_22); -lean_closure_set(x_55, 13, x_23); -lean_closure_set(x_55, 14, x_24); -lean_closure_set(x_55, 15, x_27); -lean_closure_set(x_55, 16, x_52); -x_56 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_56, 0, x_53); -lean_closure_set(x_56, 1, x_55); -x_57 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_10, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28); -return x_57; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_dec(x_130); +lean_dec(x_21); +x_146 = lean_ctor_get(x_140, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_140, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_148 = x_140; +} else { + lean_dec_ref(x_140); + x_148 = lean_box(0); +} +if (lean_is_scalar(x_148)) { + x_149 = lean_alloc_ctor(1, 2, 0); +} else { + x_149 = x_148; +} +lean_ctor_set(x_149, 0, x_146); +lean_ctor_set(x_149, 1, x_147); +return x_149; } -block_45: +} +else { -uint8_t x_30; lean_object* x_31; +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_dec(x_130); +lean_dec(x_128); +lean_dec(x_124); +lean_dec(x_117); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); lean_dec(x_29); -x_30 = 1; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_31 = l_Lean_Meta_intro1Core(x_10, x_30, x_5, x_6, x_7, x_8, x_28); -if (lean_obj_tag(x_31) == 0) +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_1); +x_150 = lean_ctor_get(x_131, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_131, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_152 = x_131; +} else { + lean_dec_ref(x_131); + x_152 = lean_box(0); +} +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 2, 0); +} else { + x_153 = x_152; +} +lean_ctor_set(x_153, 0, x_150); +lean_ctor_set(x_153, 1, x_151); +return x_153; +} +} +else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_ctor_get(x_32, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +lean_dec(x_117); lean_dec(x_32); -lean_inc(x_34); -x_36 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); -lean_closure_set(x_36, 0, x_34); -x_37 = lean_box(x_16); -lean_inc(x_35); -x_38 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 25, 16); -lean_closure_set(x_38, 0, x_35); -lean_closure_set(x_38, 1, x_11); -lean_closure_set(x_38, 2, x_12); -lean_closure_set(x_38, 3, x_13); -lean_closure_set(x_38, 4, x_14); -lean_closure_set(x_38, 5, x_15); -lean_closure_set(x_38, 6, x_37); -lean_closure_set(x_38, 7, x_17); -lean_closure_set(x_38, 8, x_18); -lean_closure_set(x_38, 9, x_19); -lean_closure_set(x_38, 10, x_20); -lean_closure_set(x_38, 11, x_21); -lean_closure_set(x_38, 12, x_22); -lean_closure_set(x_38, 13, x_23); -lean_closure_set(x_38, 14, x_24); -lean_closure_set(x_38, 15, x_34); -x_39 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_39, 0, x_36); -lean_closure_set(x_39, 1, x_38); -x_40 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_35, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_33); -return x_40; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_1); +x_154 = lean_ctor_get(x_123, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_123, 1); +lean_inc(x_155); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_156 = x_123; +} else { + lean_dec_ref(x_123); + x_156 = lean_box(0); +} +if (lean_is_scalar(x_156)) { + x_157 = lean_alloc_ctor(1, 2, 0); +} else { + x_157 = x_156; +} +lean_ctor_set(x_157, 0, x_154); +lean_ctor_set(x_157, 1, x_155); +return x_157; +} } else { -uint8_t x_41; -lean_dec(x_24); -lean_dec(x_23); +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec(x_117); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); lean_dec(x_22); lean_dec(x_21); +lean_dec(x_1); +x_158 = lean_ctor_get(x_119, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_119, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_160 = x_119; +} else { + lean_dec_ref(x_119); + x_160 = lean_box(0); +} +if (lean_is_scalar(x_160)) { + x_161 = lean_alloc_ctor(1, 2, 0); +} else { + x_161 = x_160; +} +lean_ctor_set(x_161, 0, x_158); +lean_ctor_set(x_161, 1, x_159); +return x_161; +} +} +} +} +else +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_22); +x_168 = lean_ctor_get(x_35, 1); +lean_inc(x_168); +lean_dec(x_35); +x_169 = l_Lean_LocalDecl_userName(x_25); +x_170 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_169, x_31, x_32, x_168); +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_170, 1); +lean_inc(x_172); +lean_dec(x_170); +x_173 = l_Lean_Expr_fvar___override(x_21); +x_174 = l_Lean_MVarId_assert(x_1, x_171, x_34, x_173, x_29, x_30, x_31, x_32, x_172); +if (lean_obj_tag(x_174) == 0) +{ +uint8_t x_175; +x_175 = !lean_is_exclusive(x_174); +if (x_175 == 0) +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_176 = lean_ctor_get(x_174, 0); +x_177 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_2); +lean_ctor_set(x_177, 2, x_3); +lean_ctor_set(x_177, 3, x_4); +lean_ctor_set(x_177, 4, x_5); +lean_ctor_set(x_177, 5, x_6); +lean_ctor_set(x_177, 6, x_8); +lean_ctor_set(x_177, 7, x_9); +lean_ctor_set(x_177, 8, x_10); +lean_ctor_set(x_177, 9, x_11); +lean_ctor_set(x_177, 10, x_12); +lean_ctor_set(x_177, 11, x_13); +lean_ctor_set(x_177, 12, x_14); +lean_ctor_set(x_177, 13, x_15); +lean_ctor_set(x_177, 14, x_16); +lean_ctor_set(x_177, 15, x_17); +lean_ctor_set(x_177, 16, x_18); +lean_ctor_set(x_177, 17, x_19); +lean_ctor_set(x_177, 18, x_20); +lean_ctor_set_uint8(x_177, sizeof(void*)*19, x_7); +x_178 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_174, 0, x_178); +return x_174; +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_179 = lean_ctor_get(x_174, 0); +x_180 = lean_ctor_get(x_174, 1); +lean_inc(x_180); +lean_inc(x_179); +lean_dec(x_174); +x_181 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_2); +lean_ctor_set(x_181, 2, x_3); +lean_ctor_set(x_181, 3, x_4); +lean_ctor_set(x_181, 4, x_5); +lean_ctor_set(x_181, 5, x_6); +lean_ctor_set(x_181, 6, x_8); +lean_ctor_set(x_181, 7, x_9); +lean_ctor_set(x_181, 8, x_10); +lean_ctor_set(x_181, 9, x_11); +lean_ctor_set(x_181, 10, x_12); +lean_ctor_set(x_181, 11, x_13); +lean_ctor_set(x_181, 12, x_14); +lean_ctor_set(x_181, 13, x_15); +lean_ctor_set(x_181, 14, x_16); +lean_ctor_set(x_181, 15, x_17); +lean_ctor_set(x_181, 16, x_18); +lean_ctor_set(x_181, 17, x_19); +lean_ctor_set(x_181, 18, x_20); +lean_ctor_set_uint8(x_181, sizeof(void*)*19, x_7); +x_182 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_182, 0, x_181); +x_183 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_180); +return x_183; +} +} +else +{ +uint8_t x_184; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_41 = !lean_is_exclusive(x_31); -if (x_41 == 0) +x_184 = !lean_is_exclusive(x_174); +if (x_184 == 0) { -return x_31; +return x_174; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_31, 0); -x_43 = lean_ctor_get(x_31, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_31); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_185 = lean_ctor_get(x_174, 0); +x_186 = lean_ctor_get(x_174, 1); +lean_inc(x_186); +lean_inc(x_185); +lean_dec(x_174); +x_187 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_187, 0, x_185); +lean_ctor_set(x_187, 1, x_186); +return x_187; } } } } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_77; -x_58 = lean_ctor_get(x_25, 0); -x_59 = lean_ctor_get(x_25, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_25); -x_77 = l_Lean_Expr_isArrow(x_58); -if (x_77 == 0) -{ -uint8_t x_78; -x_78 = l_Lean_Expr_isLet(x_58); -if (x_78 == 0) -{ -uint8_t x_79; -x_79 = l_Lean_Expr_isForall(x_58); -lean_dec(x_58); -if (x_79 == 0) -{ -lean_object* x_80; lean_object* x_81; -lean_dec(x_24); -lean_dec(x_23); +uint8_t x_188; +lean_dec(x_34); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_80 = lean_box(0); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_59); -return x_81; +lean_dec(x_1); +x_188 = !lean_is_exclusive(x_35); +if (x_188 == 0) +{ +return x_35; } else { -lean_object* x_82; -x_82 = lean_box(0); -x_60 = x_82; -goto block_76; +lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_189 = lean_ctor_get(x_35, 0); +x_190 = lean_ctor_get(x_35, 1); +lean_inc(x_190); +lean_inc(x_189); +lean_dec(x_35); +x_191 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_190); +return x_191; } } -else -{ -lean_object* x_83; -lean_dec(x_58); -x_83 = lean_box(0); -x_60 = x_83; -goto block_76; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_84 = l_Lean_Expr_bindingDomain_x21(x_58); -lean_inc(x_84); -x_85 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed), 9, 1); -lean_closure_set(x_85, 0, x_84); -x_86 = lean_box(x_16); -lean_inc(x_10); -x_87 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 26, 17); -lean_closure_set(x_87, 0, x_10); -lean_closure_set(x_87, 1, x_11); -lean_closure_set(x_87, 2, x_12); -lean_closure_set(x_87, 3, x_13); -lean_closure_set(x_87, 4, x_14); -lean_closure_set(x_87, 5, x_15); -lean_closure_set(x_87, 6, x_86); -lean_closure_set(x_87, 7, x_17); -lean_closure_set(x_87, 8, x_18); -lean_closure_set(x_87, 9, x_19); -lean_closure_set(x_87, 10, x_20); -lean_closure_set(x_87, 11, x_21); -lean_closure_set(x_87, 12, x_22); -lean_closure_set(x_87, 13, x_23); -lean_closure_set(x_87, 14, x_24); -lean_closure_set(x_87, 15, x_58); -lean_closure_set(x_87, 16, x_84); -x_88 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_88, 0, x_85); -lean_closure_set(x_88, 1, x_87); -x_89 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_10, x_88, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_59); -return x_89; +lean_object* x_10; +x_10 = l_Lean_Meta_isProp(x_1, x_5, x_6, x_7, x_8, x_9); +return x_10; +} } -block_76: +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_61; lean_object* x_62; -lean_dec(x_60); -x_61 = 1; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_62 = l_Lean_Meta_intro1Core(x_10, x_61, x_5, x_6, x_7, x_8, x_59); -if (lean_obj_tag(x_62) == 0) +uint8_t x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; +x_11 = 0; +x_12 = 1; +x_13 = 1; +x_14 = l_Lean_Meta_mkLambdaFVars(x_1, x_2, x_11, x_12, x_11, x_13, x_6, x_7, x_8, x_9, x_10); +return x_14; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1() { +_start: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_ctor_get(x_63, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_63, 1); -lean_inc(x_66); -lean_dec(x_63); -lean_inc(x_65); -x_67 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); -lean_closure_set(x_67, 0, x_65); -x_68 = lean_box(x_16); -lean_inc(x_66); -x_69 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 25, 16); -lean_closure_set(x_69, 0, x_66); -lean_closure_set(x_69, 1, x_11); -lean_closure_set(x_69, 2, x_12); -lean_closure_set(x_69, 3, x_13); -lean_closure_set(x_69, 4, x_14); -lean_closure_set(x_69, 5, x_15); -lean_closure_set(x_69, 6, x_68); -lean_closure_set(x_69, 7, x_17); -lean_closure_set(x_69, 8, x_18); -lean_closure_set(x_69, 9, x_19); -lean_closure_set(x_69, 10, x_20); -lean_closure_set(x_69, 11, x_21); -lean_closure_set(x_69, 12, x_22); -lean_closure_set(x_69, 13, x_23); -lean_closure_set(x_69, 14, x_24); -lean_closure_set(x_69, 15, x_65); -x_70 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_70, 0, x_67); -lean_closure_set(x_70, 1, x_69); -x_71 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_66, x_70, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_64); -return x_71; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("intro_with_eq", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31, lean_object* x_32, lean_object* x_33, lean_object* x_34, lean_object* x_35, lean_object* x_36) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_37; uint8_t x_38; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +x_37 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_28, x_29, x_30, x_31, x_32, x_33, x_34, x_35, x_36); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +x_40 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_40, 0, x_3); +lean_ctor_set(x_40, 1, x_4); +lean_ctor_set(x_40, 2, x_5); +lean_ctor_set(x_40, 3, x_6); +lean_ctor_set(x_40, 4, x_7); +lean_ctor_set(x_40, 5, x_8); +lean_ctor_set(x_40, 6, x_10); +lean_ctor_set(x_40, 7, x_11); +lean_ctor_set(x_40, 8, x_12); +lean_ctor_set(x_40, 9, x_13); +lean_ctor_set(x_40, 10, x_14); +lean_ctor_set(x_40, 11, x_15); +lean_ctor_set(x_40, 12, x_16); +lean_ctor_set(x_40, 13, x_17); +lean_ctor_set(x_40, 14, x_18); +lean_ctor_set(x_40, 15, x_19); +lean_ctor_set(x_40, 16, x_20); +lean_ctor_set(x_40, 17, x_21); +lean_ctor_set(x_40, 18, x_22); +lean_ctor_set_uint8(x_40, sizeof(void*)*19, x_9); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_23); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_37, 0, x_41); +return x_37; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_37, 1); +lean_inc(x_42); +lean_dec(x_37); +x_43 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_4); +lean_ctor_set(x_43, 2, x_5); +lean_ctor_set(x_43, 3, x_6); +lean_ctor_set(x_43, 4, x_7); +lean_ctor_set(x_43, 5, x_8); +lean_ctor_set(x_43, 6, x_10); +lean_ctor_set(x_43, 7, x_11); +lean_ctor_set(x_43, 8, x_12); +lean_ctor_set(x_43, 9, x_13); +lean_ctor_set(x_43, 10, x_14); +lean_ctor_set(x_43, 11, x_15); +lean_ctor_set(x_43, 12, x_16); +lean_ctor_set(x_43, 13, x_17); +lean_ctor_set(x_43, 14, x_18); +lean_ctor_set(x_43, 15, x_19); +lean_ctor_set(x_43, 16, x_20); +lean_ctor_set(x_43, 17, x_21); +lean_ctor_set(x_43, 18, x_22); +lean_ctor_set_uint8(x_43, sizeof(void*)*19, x_9); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_23); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_42); +return x_45; +} +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_46 = lean_ctor_get(x_1, 0); +x_47 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4; +lean_inc(x_24); +x_48 = l_Lean_Expr_const___override(x_47, x_24); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_28); +lean_ctor_set(x_49, 1, x_24); +lean_inc(x_46); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_46); +lean_ctor_set(x_50, 1, x_49); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_25); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_26); +lean_ctor_set(x_52, 1, x_51); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_27); +lean_ctor_set(x_53, 1, x_52); +x_54 = lean_array_mk(x_53); +x_55 = l_Lean_mkAppN(x_48, x_54); +lean_dec(x_54); +x_56 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_2, x_55, x_29, x_30, x_31, x_32, x_33, x_34, x_35, x_36); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_56, 0); +lean_dec(x_58); +x_59 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_59, 0, x_3); +lean_ctor_set(x_59, 1, x_4); +lean_ctor_set(x_59, 2, x_5); +lean_ctor_set(x_59, 3, x_6); +lean_ctor_set(x_59, 4, x_7); +lean_ctor_set(x_59, 5, x_8); +lean_ctor_set(x_59, 6, x_10); +lean_ctor_set(x_59, 7, x_11); +lean_ctor_set(x_59, 8, x_12); +lean_ctor_set(x_59, 9, x_13); +lean_ctor_set(x_59, 10, x_14); +lean_ctor_set(x_59, 11, x_15); +lean_ctor_set(x_59, 12, x_16); +lean_ctor_set(x_59, 13, x_17); +lean_ctor_set(x_59, 14, x_18); +lean_ctor_set(x_59, 15, x_19); +lean_ctor_set(x_59, 16, x_20); +lean_ctor_set(x_59, 17, x_21); +lean_ctor_set(x_59, 18, x_22); +lean_ctor_set_uint8(x_59, sizeof(void*)*19, x_9); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_23); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_56, 0, x_60); +return x_56; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_61 = lean_ctor_get(x_56, 1); +lean_inc(x_61); +lean_dec(x_56); +x_62 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_62, 0, x_3); +lean_ctor_set(x_62, 1, x_4); +lean_ctor_set(x_62, 2, x_5); +lean_ctor_set(x_62, 3, x_6); +lean_ctor_set(x_62, 4, x_7); +lean_ctor_set(x_62, 5, x_8); +lean_ctor_set(x_62, 6, x_10); +lean_ctor_set(x_62, 7, x_11); +lean_ctor_set(x_62, 8, x_12); +lean_ctor_set(x_62, 9, x_13); +lean_ctor_set(x_62, 10, x_14); +lean_ctor_set(x_62, 11, x_15); +lean_ctor_set(x_62, 12, x_16); +lean_ctor_set(x_62, 13, x_17); +lean_ctor_set(x_62, 14, x_18); +lean_ctor_set(x_62, 15, x_19); +lean_ctor_set(x_62, 16, x_20); +lean_ctor_set(x_62, 17, x_21); +lean_ctor_set(x_62, 18, x_22); +lean_ctor_set_uint8(x_62, sizeof(void*)*19, x_9); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_23); +lean_ctor_set(x_63, 1, x_62); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_61); +return x_64; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, uint8_t x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31) { +_start: +{ +if (x_23 == 0) +{ +uint8_t x_32; lean_object* x_33; +lean_dec(x_26); +lean_dec(x_25); lean_dec(x_24); -lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); +x_32 = 1; +x_33 = l_Lean_Meta_intro1Core(x_1, x_32, x_27, x_28, x_29, x_30, x_31); +if (lean_obj_tag(x_33) == 0) +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; uint8_t x_36; +x_35 = lean_ctor_get(x_33, 0); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 1); +x_38 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_2); +lean_ctor_set(x_38, 2, x_3); +lean_ctor_set(x_38, 3, x_4); +lean_ctor_set(x_38, 4, x_5); +lean_ctor_set(x_38, 5, x_6); +lean_ctor_set(x_38, 6, x_8); +lean_ctor_set(x_38, 7, x_9); +lean_ctor_set(x_38, 8, x_10); +lean_ctor_set(x_38, 9, x_11); +lean_ctor_set(x_38, 10, x_12); +lean_ctor_set(x_38, 11, x_13); +lean_ctor_set(x_38, 12, x_14); +lean_ctor_set(x_38, 13, x_15); +lean_ctor_set(x_38, 14, x_16); +lean_ctor_set(x_38, 15, x_17); +lean_ctor_set(x_38, 16, x_18); +lean_ctor_set(x_38, 17, x_19); +lean_ctor_set(x_38, 18, x_20); +lean_ctor_set_uint8(x_38, sizeof(void*)*19, x_7); +lean_ctor_set_tag(x_35, 3); +lean_ctor_set(x_35, 1, x_38); +return x_33; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_35, 0); +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_35); +x_41 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_2); +lean_ctor_set(x_41, 2, x_3); +lean_ctor_set(x_41, 3, x_4); +lean_ctor_set(x_41, 4, x_5); +lean_ctor_set(x_41, 5, x_6); +lean_ctor_set(x_41, 6, x_8); +lean_ctor_set(x_41, 7, x_9); +lean_ctor_set(x_41, 8, x_10); +lean_ctor_set(x_41, 9, x_11); +lean_ctor_set(x_41, 10, x_12); +lean_ctor_set(x_41, 11, x_13); +lean_ctor_set(x_41, 12, x_14); +lean_ctor_set(x_41, 13, x_15); +lean_ctor_set(x_41, 14, x_16); +lean_ctor_set(x_41, 15, x_17); +lean_ctor_set(x_41, 16, x_18); +lean_ctor_set(x_41, 17, x_19); +lean_ctor_set(x_41, 18, x_20); +lean_ctor_set_uint8(x_41, sizeof(void*)*19, x_7); +x_42 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_33, 0, x_42); +return x_33; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_43 = lean_ctor_get(x_33, 0); +x_44 = lean_ctor_get(x_33, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_33); +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_43)) { + lean_ctor_release(x_43, 0); + lean_ctor_release(x_43, 1); + x_47 = x_43; +} else { + lean_dec_ref(x_43); + x_47 = lean_box(0); +} +x_48 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_2); +lean_ctor_set(x_48, 2, x_3); +lean_ctor_set(x_48, 3, x_4); +lean_ctor_set(x_48, 4, x_5); +lean_ctor_set(x_48, 5, x_6); +lean_ctor_set(x_48, 6, x_8); +lean_ctor_set(x_48, 7, x_9); +lean_ctor_set(x_48, 8, x_10); +lean_ctor_set(x_48, 9, x_11); +lean_ctor_set(x_48, 10, x_12); +lean_ctor_set(x_48, 11, x_13); +lean_ctor_set(x_48, 12, x_14); +lean_ctor_set(x_48, 13, x_15); +lean_ctor_set(x_48, 14, x_16); +lean_ctor_set(x_48, 15, x_17); +lean_ctor_set(x_48, 16, x_18); +lean_ctor_set(x_48, 17, x_19); +lean_ctor_set(x_48, 18, x_20); +lean_ctor_set_uint8(x_48, sizeof(void*)*19, x_7); +if (lean_is_scalar(x_47)) { + x_49 = lean_alloc_ctor(3, 2, 0); +} else { + x_49 = x_47; + lean_ctor_set_tag(x_49, 3); +} +lean_ctor_set(x_49, 0, x_45); +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_44); +return x_50; +} +} +else +{ +uint8_t x_51; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_72 = lean_ctor_get(x_62, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_62, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_74 = x_62; -} else { - lean_dec_ref(x_62); - x_74 = lean_box(0); +x_51 = !lean_is_exclusive(x_33); +if (x_51 == 0) +{ +return x_33; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_33, 0); +x_53 = lean_ctor_get(x_33, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_33); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(1, 2, 0); -} else { - x_75 = x_74; } -lean_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; } +else +{ +lean_object* x_55; +lean_inc(x_1); +x_55 = l_Lean_MVarId_getTag(x_1, x_27, x_28, x_29, x_30, x_31); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = l_Lean_Expr_bindingBody_x21(x_21); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_25); +lean_inc(x_22); +x_59 = l_Lean_Meta_Grind_simp(x_22, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_57); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec(x_59); +x_62 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1(x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_61); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_27, 2); +lean_inc(x_65); +x_66 = l_Lean_Expr_bindingName_x21(x_21); +x_67 = lean_ctor_get(x_60, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_60, 1); +lean_inc(x_68); +lean_dec(x_60); +x_69 = l_Lean_Expr_bindingInfo_x21(x_21); +x_70 = lean_unbox(x_69); +lean_dec(x_69); +x_71 = 0; +lean_inc(x_67); +lean_inc(x_63); +x_72 = l_Lean_LocalContext_mkLocalDecl(x_65, x_63, x_66, x_67, x_70, x_71); +x_73 = l_Lean_Meta_getLocalInstances(x_27, x_28, x_29, x_30, x_64); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = 2; +x_77 = lean_unsigned_to_nat(0u); +lean_inc(x_58); +x_78 = l_Lean_Meta_mkFreshExprMVarAt(x_72, x_74, x_58, x_76, x_56, x_77, x_27, x_28, x_29, x_30, x_75); +x_79 = !lean_is_exclusive(x_78); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_80 = lean_ctor_get(x_78, 0); +x_81 = lean_ctor_get(x_78, 1); +x_82 = l_Lean_Expr_mvarId_x21(x_80); +lean_inc(x_63); +x_83 = l_Lean_Expr_fvar___override(x_63); +x_84 = lean_box(0); +lean_ctor_set_tag(x_78, 1); +lean_ctor_set(x_78, 1, x_84); +lean_ctor_set(x_78, 0, x_83); +x_85 = lean_array_mk(x_78); +x_86 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7___boxed), 10, 2); +lean_closure_set(x_86, 0, x_85); +lean_closure_set(x_86, 1, x_80); +x_87 = lean_box(x_7); +lean_inc(x_82); +x_88 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___boxed), 36, 27); +lean_closure_set(x_88, 0, x_68); +lean_closure_set(x_88, 1, x_1); +lean_closure_set(x_88, 2, x_82); +lean_closure_set(x_88, 3, x_2); +lean_closure_set(x_88, 4, x_3); +lean_closure_set(x_88, 5, x_4); +lean_closure_set(x_88, 6, x_5); +lean_closure_set(x_88, 7, x_6); +lean_closure_set(x_88, 8, x_87); +lean_closure_set(x_88, 9, x_8); +lean_closure_set(x_88, 10, x_9); +lean_closure_set(x_88, 11, x_10); +lean_closure_set(x_88, 12, x_11); +lean_closure_set(x_88, 13, x_12); +lean_closure_set(x_88, 14, x_13); +lean_closure_set(x_88, 15, x_14); +lean_closure_set(x_88, 16, x_15); +lean_closure_set(x_88, 17, x_16); +lean_closure_set(x_88, 18, x_17); +lean_closure_set(x_88, 19, x_18); +lean_closure_set(x_88, 20, x_19); +lean_closure_set(x_88, 21, x_20); +lean_closure_set(x_88, 22, x_63); +lean_closure_set(x_88, 23, x_84); +lean_closure_set(x_88, 24, x_58); +lean_closure_set(x_88, 25, x_67); +lean_closure_set(x_88, 26, x_22); +x_89 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_89, 0, x_86); +lean_closure_set(x_89, 1, x_88); +x_90 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_82, x_89, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_81); +return x_90; } +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_91 = lean_ctor_get(x_78, 0); +x_92 = lean_ctor_get(x_78, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_78); +x_93 = l_Lean_Expr_mvarId_x21(x_91); +lean_inc(x_63); +x_94 = l_Lean_Expr_fvar___override(x_63); +x_95 = lean_box(0); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_array_mk(x_96); +x_98 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7___boxed), 10, 2); +lean_closure_set(x_98, 0, x_97); +lean_closure_set(x_98, 1, x_91); +x_99 = lean_box(x_7); +lean_inc(x_93); +x_100 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___boxed), 36, 27); +lean_closure_set(x_100, 0, x_68); +lean_closure_set(x_100, 1, x_1); +lean_closure_set(x_100, 2, x_93); +lean_closure_set(x_100, 3, x_2); +lean_closure_set(x_100, 4, x_3); +lean_closure_set(x_100, 5, x_4); +lean_closure_set(x_100, 6, x_5); +lean_closure_set(x_100, 7, x_6); +lean_closure_set(x_100, 8, x_99); +lean_closure_set(x_100, 9, x_8); +lean_closure_set(x_100, 10, x_9); +lean_closure_set(x_100, 11, x_10); +lean_closure_set(x_100, 12, x_11); +lean_closure_set(x_100, 13, x_12); +lean_closure_set(x_100, 14, x_13); +lean_closure_set(x_100, 15, x_14); +lean_closure_set(x_100, 16, x_15); +lean_closure_set(x_100, 17, x_16); +lean_closure_set(x_100, 18, x_17); +lean_closure_set(x_100, 19, x_18); +lean_closure_set(x_100, 20, x_19); +lean_closure_set(x_100, 21, x_20); +lean_closure_set(x_100, 22, x_63); +lean_closure_set(x_100, 23, x_95); +lean_closure_set(x_100, 24, x_58); +lean_closure_set(x_100, 25, x_67); +lean_closure_set(x_100, 26, x_22); +x_101 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_101, 0, x_98); +lean_closure_set(x_101, 1, x_100); +x_102 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_93, x_101, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_92); +return x_102; } } else { -uint8_t x_90; +uint8_t x_103; +lean_dec(x_58); +lean_dec(x_56); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); lean_dec(x_24); -lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_90 = !lean_is_exclusive(x_25); -if (x_90 == 0) +lean_dec(x_1); +x_103 = !lean_is_exclusive(x_59); +if (x_103 == 0) { -return x_25; +return x_59; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_25, 0); -x_92 = lean_ctor_get(x_25, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_25); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; -} -} -} +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_59, 0); +x_105 = lean_ctor_get(x_59, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_59); +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; } -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; -x_7 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +uint8_t x_107; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_7; -} +x_107 = !lean_is_exclusive(x_55); +if (x_107 == 0) +{ +return x_55; } -LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_55, 0); +x_109 = lean_ctor_get(x_55, 1); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_55); +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_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; -x_11 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_1, 3); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 4); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 5); +lean_inc(x_16); +x_17 = lean_ctor_get_uint8(x_1, sizeof(void*)*19); +x_18 = lean_ctor_get(x_1, 6); +lean_inc(x_18); +x_19 = lean_ctor_get(x_1, 7); +lean_inc(x_19); +x_20 = lean_ctor_get(x_1, 8); +lean_inc(x_20); +x_21 = lean_ctor_get(x_1, 9); +lean_inc(x_21); +x_22 = lean_ctor_get(x_1, 10); +lean_inc(x_22); +x_23 = lean_ctor_get(x_1, 11); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 12); +lean_inc(x_24); +x_25 = lean_ctor_get(x_1, 13); +lean_inc(x_25); +x_26 = lean_ctor_get(x_1, 14); +lean_inc(x_26); +x_27 = lean_ctor_get(x_1, 15); +lean_inc(x_27); +x_28 = lean_ctor_get(x_1, 16); +lean_inc(x_28); +x_29 = lean_ctor_get(x_1, 17); +lean_inc(x_29); +x_30 = lean_ctor_get(x_1, 18); +lean_inc(x_30); +lean_dec(x_1); +lean_inc(x_11); +x_31 = l_Lean_MVarId_getType(x_11, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +x_35 = l_Lean_Expr_isArrow(x_33); +if (x_35 == 0) +{ +uint8_t x_36; lean_object* x_37; +x_36 = l_Lean_Expr_isLet(x_33); +if (x_36 == 0) +{ +uint8_t x_55; +x_55 = l_Lean_Expr_isForall(x_33); +if (x_55 == 0) +{ +uint8_t x_56; +x_56 = l_Lean_Expr_isLetFun(x_33); +if (x_56 == 0) +{ +lean_object* x_57; +lean_dec(x_33); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -1950,78 +2472,233 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_11; +lean_dec(x_2); +x_57 = lean_box(0); +lean_ctor_set(x_31, 0, x_57); +return x_31; } +else +{ +lean_object* x_58; +lean_free_object(x_31); +x_58 = lean_box(0); +x_37 = x_58; +goto block_54; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +} +else { -lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_59; +lean_free_object(x_31); +x_59 = lean_box(0); +x_37 = x_59; +goto block_54; +} +} +else +{ +lean_object* x_60; +lean_free_object(x_31); +x_60 = lean_box(0); +x_37 = x_60; +goto block_54; +} +block_54: +{ +uint8_t x_38; lean_object* x_39; +lean_dec(x_37); +x_38 = 1; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_39 = l_Lean_Meta_intro1Core(x_11, x_38, x_6, x_7, x_8, x_9, x_34); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +lean_inc(x_42); +x_44 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); +lean_closure_set(x_44, 0, x_42); +x_45 = lean_box(x_17); +x_46 = lean_box(x_36); +lean_inc(x_43); +x_47 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 33, 24); +lean_closure_set(x_47, 0, x_43); +lean_closure_set(x_47, 1, x_12); +lean_closure_set(x_47, 2, x_13); +lean_closure_set(x_47, 3, x_14); +lean_closure_set(x_47, 4, x_15); +lean_closure_set(x_47, 5, x_16); +lean_closure_set(x_47, 6, x_45); +lean_closure_set(x_47, 7, x_18); +lean_closure_set(x_47, 8, x_19); +lean_closure_set(x_47, 9, x_20); +lean_closure_set(x_47, 10, x_21); +lean_closure_set(x_47, 11, x_22); +lean_closure_set(x_47, 12, x_23); +lean_closure_set(x_47, 13, x_24); +lean_closure_set(x_47, 14, x_25); +lean_closure_set(x_47, 15, x_26); +lean_closure_set(x_47, 16, x_27); +lean_closure_set(x_47, 17, x_28); +lean_closure_set(x_47, 18, x_29); +lean_closure_set(x_47, 19, x_30); +lean_closure_set(x_47, 20, x_42); +lean_closure_set(x_47, 21, x_2); +lean_closure_set(x_47, 22, x_46); +lean_closure_set(x_47, 23, x_33); +x_48 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_48, 0, x_44); +lean_closure_set(x_48, 1, x_47); +x_49 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_43, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_41); +return x_49; +} +else +{ +uint8_t x_50; +lean_dec(x_33); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_10; -} +x_50 = !lean_is_exclusive(x_39); +if (x_50 == 0) +{ +return x_39; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -lean_object* x_24 = _args[23]; -lean_object* x_25 = _args[24]; -_start: +else { -uint8_t x_26; lean_object* x_27; -x_26 = lean_unbox(x_7); -lean_dec(x_7); -x_27 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_26, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -return x_27; +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_39, 0); +x_52 = lean_ctor_get(x_39, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_39); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +} +} +else { -lean_object* x_10; -x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_3); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_free_object(x_31); lean_dec(x_2); -return x_10; +x_61 = l_Lean_Expr_bindingDomain_x21(x_33); +lean_inc(x_61); +x_62 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 9, 1); +lean_closure_set(x_62, 0, x_61); +x_63 = lean_box(x_17); +lean_inc(x_11); +x_64 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9___boxed), 31, 22); +lean_closure_set(x_64, 0, x_11); +lean_closure_set(x_64, 1, x_12); +lean_closure_set(x_64, 2, x_13); +lean_closure_set(x_64, 3, x_14); +lean_closure_set(x_64, 4, x_15); +lean_closure_set(x_64, 5, x_16); +lean_closure_set(x_64, 6, x_63); +lean_closure_set(x_64, 7, x_18); +lean_closure_set(x_64, 8, x_19); +lean_closure_set(x_64, 9, x_20); +lean_closure_set(x_64, 10, x_21); +lean_closure_set(x_64, 11, x_22); +lean_closure_set(x_64, 12, x_23); +lean_closure_set(x_64, 13, x_24); +lean_closure_set(x_64, 14, x_25); +lean_closure_set(x_64, 15, x_26); +lean_closure_set(x_64, 16, x_27); +lean_closure_set(x_64, 17, x_28); +lean_closure_set(x_64, 18, x_29); +lean_closure_set(x_64, 19, x_30); +lean_closure_set(x_64, 20, x_33); +lean_closure_set(x_64, 21, x_61); +x_65 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_65, 0, x_62); +lean_closure_set(x_65, 1, x_64); +x_66 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_65, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); +return x_66; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = lean_ctor_get(x_31, 0); +x_68 = lean_ctor_get(x_31, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_31); +x_69 = l_Lean_Expr_isArrow(x_67); +if (x_69 == 0) +{ +uint8_t x_70; lean_object* x_71; +x_70 = l_Lean_Expr_isLet(x_67); +if (x_70 == 0) +{ +uint8_t x_89; +x_89 = l_Lean_Expr_isForall(x_67); +if (x_89 == 0) +{ +uint8_t x_90; +x_90 = l_Lean_Expr_isLetFun(x_67); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; +lean_dec(x_67); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2029,12 +2706,340 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -return x_11; +lean_dec(x_2); +x_91 = lean_box(0); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_68); +return x_92; } +else +{ +lean_object* x_93; +x_93 = lean_box(0); +x_71 = x_93; +goto block_88; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; +} +else +{ +lean_object* x_94; +x_94 = lean_box(0); +x_71 = x_94; +goto block_88; +} +} +else +{ +lean_object* x_95; +x_95 = lean_box(0); +x_71 = x_95; +goto block_88; +} +block_88: +{ +uint8_t x_72; lean_object* x_73; +lean_dec(x_71); +x_72 = 1; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_73 = l_Lean_Meta_intro1Core(x_11, x_72, x_6, x_7, x_8, x_9, x_68); +if (lean_obj_tag(x_73) == 0) +{ +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_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = lean_ctor_get(x_74, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_74, 1); +lean_inc(x_77); +lean_dec(x_74); +lean_inc(x_76); +x_78 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed), 9, 1); +lean_closure_set(x_78, 0, x_76); +x_79 = lean_box(x_17); +x_80 = lean_box(x_70); +lean_inc(x_77); +x_81 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed), 33, 24); +lean_closure_set(x_81, 0, x_77); +lean_closure_set(x_81, 1, x_12); +lean_closure_set(x_81, 2, x_13); +lean_closure_set(x_81, 3, x_14); +lean_closure_set(x_81, 4, x_15); +lean_closure_set(x_81, 5, x_16); +lean_closure_set(x_81, 6, x_79); +lean_closure_set(x_81, 7, x_18); +lean_closure_set(x_81, 8, x_19); +lean_closure_set(x_81, 9, x_20); +lean_closure_set(x_81, 10, x_21); +lean_closure_set(x_81, 11, x_22); +lean_closure_set(x_81, 12, x_23); +lean_closure_set(x_81, 13, x_24); +lean_closure_set(x_81, 14, x_25); +lean_closure_set(x_81, 15, x_26); +lean_closure_set(x_81, 16, x_27); +lean_closure_set(x_81, 17, x_28); +lean_closure_set(x_81, 18, x_29); +lean_closure_set(x_81, 19, x_30); +lean_closure_set(x_81, 20, x_76); +lean_closure_set(x_81, 21, x_2); +lean_closure_set(x_81, 22, x_80); +lean_closure_set(x_81, 23, x_67); +x_82 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_82, 0, x_78); +lean_closure_set(x_82, 1, x_81); +x_83 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_77, x_82, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_75); +return x_83; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_67); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_84 = lean_ctor_get(x_73, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_73, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_86 = x_73; +} else { + lean_dec_ref(x_73); + x_86 = lean_box(0); +} +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(1, 2, 0); +} else { + x_87 = x_86; +} +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_85); +return x_87; +} +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +lean_dec(x_2); +x_96 = l_Lean_Expr_bindingDomain_x21(x_67); +lean_inc(x_96); +x_97 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed), 9, 1); +lean_closure_set(x_97, 0, x_96); +x_98 = lean_box(x_17); +lean_inc(x_11); +x_99 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9___boxed), 31, 22); +lean_closure_set(x_99, 0, x_11); +lean_closure_set(x_99, 1, x_12); +lean_closure_set(x_99, 2, x_13); +lean_closure_set(x_99, 3, x_14); +lean_closure_set(x_99, 4, x_15); +lean_closure_set(x_99, 5, x_16); +lean_closure_set(x_99, 6, x_98); +lean_closure_set(x_99, 7, x_18); +lean_closure_set(x_99, 8, x_19); +lean_closure_set(x_99, 9, x_20); +lean_closure_set(x_99, 10, x_21); +lean_closure_set(x_99, 11, x_22); +lean_closure_set(x_99, 12, x_23); +lean_closure_set(x_99, 13, x_24); +lean_closure_set(x_99, 14, x_25); +lean_closure_set(x_99, 15, x_26); +lean_closure_set(x_99, 16, x_27); +lean_closure_set(x_99, 17, x_28); +lean_closure_set(x_99, 18, x_29); +lean_closure_set(x_99, 19, x_30); +lean_closure_set(x_99, 20, x_67); +lean_closure_set(x_99, 21, x_96); +x_100 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_100, 0, x_97); +lean_closure_set(x_100, 1, x_99); +x_101 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_11, x_100, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_68); +return x_101; +} +} +} +else +{ +uint8_t x_102; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_102 = !lean_is_exclusive(x_31); +if (x_102 == 0) +{ +return x_31; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_31, 0); +x_104 = lean_ctor_get(x_31, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_31); +x_105 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 1, x_104); +return x_105; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_mkFreshFVarId___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_MVarId_assign___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; lean_object* x_4 = _args[3]; @@ -2065,24 +3070,49 @@ lean_object* x_28 = _args[27]; lean_object* x_29 = _args[28]; lean_object* x_30 = _args[29]; lean_object* x_31 = _args[30]; +lean_object* x_32 = _args[31]; +lean_object* x_33 = _args[32]; _start: { -uint8_t x_32; lean_object* x_33; -x_32 = lean_unbox(x_9); -lean_dec(x_9); -x_33 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_28); -lean_dec(x_27); -lean_dec(x_26); +uint8_t x_34; uint8_t x_35; lean_object* x_36; +x_34 = lean_unbox(x_7); +lean_dec(x_7); +x_35 = lean_unbox(x_23); +lean_dec(x_23); +x_36 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_34, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_35, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_33); lean_dec(x_25); lean_dec(x_24); +return x_36; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -return x_33; +return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -2109,109 +3139,102 @@ lean_object* x_23 = _args[22]; lean_object* x_24 = _args[23]; lean_object* x_25 = _args[24]; lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +lean_object* x_29 = _args[28]; +lean_object* x_30 = _args[29]; +lean_object* x_31 = _args[30]; +lean_object* x_32 = _args[31]; +lean_object* x_33 = _args[32]; +lean_object* x_34 = _args[33]; +lean_object* x_35 = _args[34]; +lean_object* x_36 = _args[35]; _start: { -uint8_t x_27; uint8_t x_28; lean_object* x_29; -x_27 = lean_unbox(x_7); -lean_dec(x_7); -x_28 = lean_unbox(x_18); -lean_dec(x_18); -x_29 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_27, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_28, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26); -lean_dec(x_16); -return x_29; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_FVarId_getType(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Expr_getAppFn(x_9); +uint8_t x_37; lean_object* x_38; +x_37 = lean_unbox(x_9); lean_dec(x_9); -if (lean_obj_tag(x_11) == 4) -{ -lean_object* x_12; lean_object* x_13; -lean_free_object(x_7); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Meta_Grind_isGrindCasesTarget(x_12, x_4, x_5, x_10); -lean_dec(x_12); -return x_13; -} -else -{ -uint8_t x_14; lean_object* x_15; -lean_dec(x_11); -x_14 = 0; -x_15 = lean_box(x_14); -lean_ctor_set(x_7, 0, x_15); -return x_7; +x_38 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_37, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_33, x_34, x_35, x_36); +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_1); +return x_38; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +lean_object* x_29 = _args[28]; +lean_object* x_30 = _args[29]; +lean_object* x_31 = _args[30]; +_start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_7, 0); -x_17 = lean_ctor_get(x_7, 1); -lean_inc(x_17); -lean_inc(x_16); +uint8_t x_32; uint8_t x_33; lean_object* x_34; +x_32 = lean_unbox(x_7); lean_dec(x_7); -x_18 = l_Lean_Expr_getAppFn(x_16); -lean_dec(x_16); -if (lean_obj_tag(x_18) == 4) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Meta_Grind_isGrindCasesTarget(x_19, x_4, x_5, x_17); -lean_dec(x_19); -return x_20; -} -else -{ -uint8_t x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_18); -x_21 = 0; -x_22 = lean_box(x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_17); -return x_23; -} +x_33 = lean_unbox(x_23); +lean_dec(x_23); +x_34 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_32, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_33, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31); +lean_dec(x_21); +return x_34; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_7); -if (x_24 == 0) +lean_object* x_7; +x_7 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_7) == 4) { -return x_7; +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +x_9 = l_Lean_Meta_Grind_isGrindCasesTarget(x_8, x_4, x_5, x_6); +lean_dec(x_8); +return x_9; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_7, 0); -x_26 = lean_ctor_get(x_7, 1); -lean_inc(x_26); -lean_inc(x_25); +uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_dec(x_7); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; -} +x_10 = 0; +x_11 = lean_box(x_10); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_6); +return x_12; } } } @@ -2223,6 +3246,8 @@ x_7 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandida lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); return x_7; } } @@ -2241,7 +3266,7 @@ uint8_t x_5; x_5 = !lean_is_exclusive(x_2); if (x_5 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_6 = lean_ctor_get(x_2, 0); x_7 = lean_ctor_get(x_2, 1); x_8 = lean_ctor_get(x_1, 1); @@ -2249,7 +3274,7 @@ x_9 = lean_ctor_get(x_1, 2); x_10 = lean_ctor_get(x_1, 3); x_11 = lean_ctor_get(x_1, 4); x_12 = lean_ctor_get(x_1, 5); -x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*19); x_14 = lean_ctor_get(x_1, 6); x_15 = lean_ctor_get(x_1, 7); x_16 = lean_ctor_get(x_1, 8); @@ -2258,6 +3283,16 @@ x_18 = lean_ctor_get(x_1, 10); x_19 = lean_ctor_get(x_1, 11); x_20 = lean_ctor_get(x_1, 12); x_21 = lean_ctor_get(x_1, 13); +x_22 = lean_ctor_get(x_1, 14); +x_23 = lean_ctor_get(x_1, 15); +x_24 = lean_ctor_get(x_1, 16); +x_25 = lean_ctor_get(x_1, 17); +x_26 = lean_ctor_get(x_1, 18); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); @@ -2271,24 +3306,29 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_22 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_22, 0, x_6); -lean_ctor_set(x_22, 1, x_8); -lean_ctor_set(x_22, 2, x_9); -lean_ctor_set(x_22, 3, x_10); -lean_ctor_set(x_22, 4, x_11); -lean_ctor_set(x_22, 5, x_12); -lean_ctor_set(x_22, 6, x_14); -lean_ctor_set(x_22, 7, x_15); -lean_ctor_set(x_22, 8, x_16); -lean_ctor_set(x_22, 9, x_17); -lean_ctor_set(x_22, 10, x_18); -lean_ctor_set(x_22, 11, x_19); -lean_ctor_set(x_22, 12, x_20); -lean_ctor_set(x_22, 13, x_21); -lean_ctor_set_uint8(x_22, sizeof(void*)*14, x_13); +x_27 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_8); +lean_ctor_set(x_27, 2, x_9); +lean_ctor_set(x_27, 3, x_10); +lean_ctor_set(x_27, 4, x_11); +lean_ctor_set(x_27, 5, x_12); +lean_ctor_set(x_27, 6, x_14); +lean_ctor_set(x_27, 7, x_15); +lean_ctor_set(x_27, 8, x_16); +lean_ctor_set(x_27, 9, x_17); +lean_ctor_set(x_27, 10, x_18); +lean_ctor_set(x_27, 11, x_19); +lean_ctor_set(x_27, 12, x_20); +lean_ctor_set(x_27, 13, x_21); +lean_ctor_set(x_27, 14, x_22); +lean_ctor_set(x_27, 15, x_23); +lean_ctor_set(x_27, 16, x_24); +lean_ctor_set(x_27, 17, x_25); +lean_ctor_set(x_27, 18, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*19, x_13); lean_ctor_set(x_2, 1, x_3); -lean_ctor_set(x_2, 0, x_22); +lean_ctor_set(x_2, 0, x_27); { lean_object* _tmp_1 = x_7; lean_object* _tmp_2 = x_2; @@ -2299,60 +3339,75 @@ goto _start; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_24 = lean_ctor_get(x_2, 0); -x_25 = lean_ctor_get(x_2, 1); -lean_inc(x_25); -lean_inc(x_24); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_29 = lean_ctor_get(x_2, 0); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_inc(x_29); lean_dec(x_2); -x_26 = lean_ctor_get(x_1, 1); -x_27 = lean_ctor_get(x_1, 2); -x_28 = lean_ctor_get(x_1, 3); -x_29 = lean_ctor_get(x_1, 4); -x_30 = lean_ctor_get(x_1, 5); -x_31 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_32 = lean_ctor_get(x_1, 6); -x_33 = lean_ctor_get(x_1, 7); -x_34 = lean_ctor_get(x_1, 8); -x_35 = lean_ctor_get(x_1, 9); -x_36 = lean_ctor_get(x_1, 10); -x_37 = lean_ctor_get(x_1, 11); -x_38 = lean_ctor_get(x_1, 12); -x_39 = lean_ctor_get(x_1, 13); +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_ctor_get(x_1, 2); +x_33 = lean_ctor_get(x_1, 3); +x_34 = lean_ctor_get(x_1, 4); +x_35 = lean_ctor_get(x_1, 5); +x_36 = lean_ctor_get_uint8(x_1, sizeof(void*)*19); +x_37 = lean_ctor_get(x_1, 6); +x_38 = lean_ctor_get(x_1, 7); +x_39 = lean_ctor_get(x_1, 8); +x_40 = lean_ctor_get(x_1, 9); +x_41 = lean_ctor_get(x_1, 10); +x_42 = lean_ctor_get(x_1, 11); +x_43 = lean_ctor_get(x_1, 12); +x_44 = lean_ctor_get(x_1, 13); +x_45 = lean_ctor_get(x_1, 14); +x_46 = lean_ctor_get(x_1, 15); +x_47 = lean_ctor_get(x_1, 16); +x_48 = lean_ctor_get(x_1, 17); +x_49 = lean_ctor_get(x_1, 18); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); lean_inc(x_39); lean_inc(x_38); lean_inc(x_37); -lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -x_40 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_40, 0, x_24); -lean_ctor_set(x_40, 1, x_26); -lean_ctor_set(x_40, 2, x_27); -lean_ctor_set(x_40, 3, x_28); -lean_ctor_set(x_40, 4, x_29); -lean_ctor_set(x_40, 5, x_30); -lean_ctor_set(x_40, 6, x_32); -lean_ctor_set(x_40, 7, x_33); -lean_ctor_set(x_40, 8, x_34); -lean_ctor_set(x_40, 9, x_35); -lean_ctor_set(x_40, 10, x_36); -lean_ctor_set(x_40, 11, x_37); -lean_ctor_set(x_40, 12, x_38); -lean_ctor_set(x_40, 13, x_39); -lean_ctor_set_uint8(x_40, sizeof(void*)*14, x_31); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_3); -x_2 = x_25; -x_3 = x_41; +lean_inc(x_31); +x_50 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_50, 0, x_29); +lean_ctor_set(x_50, 1, x_31); +lean_ctor_set(x_50, 2, x_32); +lean_ctor_set(x_50, 3, x_33); +lean_ctor_set(x_50, 4, x_34); +lean_ctor_set(x_50, 5, x_35); +lean_ctor_set(x_50, 6, x_37); +lean_ctor_set(x_50, 7, x_38); +lean_ctor_set(x_50, 8, x_39); +lean_ctor_set(x_50, 9, x_40); +lean_ctor_set(x_50, 10, x_41); +lean_ctor_set(x_50, 11, x_42); +lean_ctor_set(x_50, 12, x_43); +lean_ctor_set(x_50, 13, x_44); +lean_ctor_set(x_50, 14, x_45); +lean_ctor_set(x_50, 15, x_46); +lean_ctor_set(x_50, 16, x_47); +lean_ctor_set(x_50, 17, x_48); +lean_ctor_set(x_50, 18, x_49); +lean_ctor_set_uint8(x_50, sizeof(void*)*19, x_36); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_3); +x_2 = x_30; +x_3 = x_51; goto _start; } } @@ -2364,136 +3419,144 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_G lean_object* x_9; lean_inc(x_4); lean_inc(x_1); -x_9 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_1, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_FVarId_getType(x_1, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_9) == 0) { -lean_object* x_10; uint8_t x_11; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -x_11 = lean_unbox(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_10, x_4, x_5, x_6, x_7, x_11); lean_dec(x_10); -if (x_11 == 0) +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) { -uint8_t x_12; +uint8_t x_15; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_9, 0); -lean_dec(x_13); -x_14 = lean_box(0); -lean_ctor_set(x_9, 0, x_14); -return x_9; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = lean_box(0); +lean_ctor_set(x_12, 0, x_17); +return x_12; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_9, 1); -lean_inc(x_15); -lean_dec(x_9); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; } } else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_9, 1); -lean_inc(x_18); -lean_dec(x_9); -x_19 = l_Lean_Meta_Grind_cases(x_2, x_1, x_4, x_5, x_6, x_7, x_18); -if (lean_obj_tag(x_19) == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = l_Lean_Expr_fvar___override(x_1); +x_23 = l_Lean_Meta_Grind_cases(x_2, x_22, x_4, x_5, x_6, x_7, x_21); +if (lean_obj_tag(x_23) == 0) { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_box(0); -x_23 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_21, x_22); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_19, 0, x_24); -return x_19; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_box(0); +x_27 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_25, x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_23, 0, x_28); +return x_23; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_25 = lean_ctor_get(x_19, 0); -x_26 = lean_ctor_get(x_19, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_19); -x_27 = lean_box(0); -x_28 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_25, x_27); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_26); -return x_30; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_29 = lean_ctor_get(x_23, 0); +x_30 = lean_ctor_get(x_23, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_23); +x_31 = lean_box(0); +x_32 = l_List_mapTR_loop___at___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f___spec__1(x_3, x_29, x_31); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_30); +return x_34; } } else { -uint8_t x_31; -x_31 = !lean_is_exclusive(x_19); -if (x_31 == 0) +uint8_t x_35; +x_35 = !lean_is_exclusive(x_23); +if (x_35 == 0) { -return x_19; +return x_23; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_19, 0); -x_33 = lean_ctor_get(x_19, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_19); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_23, 0); +x_37 = lean_ctor_get(x_23, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_23); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } } } else { -uint8_t x_35; +uint8_t x_39; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_35 = !lean_is_exclusive(x_9); -if (x_35 == 0) +x_39 = !lean_is_exclusive(x_9); +if (x_39 == 0) { return x_9; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_9, 0); -x_37 = lean_ctor_get(x_9, 1); -lean_inc(x_37); -lean_inc(x_36); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_9, 0); +x_41 = lean_ctor_get(x_9, 1); +lean_inc(x_41); +lean_inc(x_40); lean_dec(x_9); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +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; } } } @@ -2538,7 +3601,7 @@ uint8_t x_8; x_8 = !lean_is_exclusive(x_1); if (x_8 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; x_9 = lean_ctor_get(x_1, 0); x_10 = lean_ctor_get(x_1, 1); x_11 = lean_ctor_get(x_1, 2); @@ -2553,16 +3616,26 @@ x_19 = lean_ctor_get(x_1, 10); x_20 = lean_ctor_get(x_1, 11); x_21 = lean_ctor_get(x_1, 12); x_22 = lean_ctor_get(x_1, 13); -x_23 = l_Lean_Meta_Grind_injection_x3f(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_23) == 0) +x_23 = lean_ctor_get(x_1, 14); +x_24 = lean_ctor_get(x_1, 15); +x_25 = lean_ctor_get(x_1, 16); +x_26 = lean_ctor_get(x_1, 17); +x_27 = lean_ctor_get(x_1, 18); +x_28 = l_Lean_Meta_Grind_injection_x3f(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -if (lean_obj_tag(x_24) == 0) +lean_object* x_29; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +if (lean_obj_tag(x_29) == 0) { -uint8_t x_25; +uint8_t x_30; lean_free_object(x_1); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -2576,93 +3649,98 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) +x_30 = !lean_is_exclusive(x_28); +if (x_30 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_23, 0, x_27); -return x_23; +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_28, 0); +lean_dec(x_31); +x_32 = lean_box(0); +lean_ctor_set(x_28, 0, x_32); +return x_28; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 1); -lean_inc(x_28); -lean_dec(x_23); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_28, 1); +lean_inc(x_33); +lean_dec(x_28); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; } } else { -uint8_t x_31; -x_31 = !lean_is_exclusive(x_23); -if (x_31 == 0) +uint8_t x_36; +x_36 = !lean_is_exclusive(x_28); +if (x_36 == 0) { -lean_object* x_32; uint8_t x_33; -x_32 = lean_ctor_get(x_23, 0); -lean_dec(x_32); -x_33 = !lean_is_exclusive(x_24); -if (x_33 == 0) +lean_object* x_37; uint8_t x_38; +x_37 = lean_ctor_get(x_28, 0); +lean_dec(x_37); +x_38 = !lean_is_exclusive(x_29); +if (x_38 == 0) { -lean_object* x_34; -x_34 = lean_ctor_get(x_24, 0); -lean_ctor_set(x_1, 0, x_34); -lean_ctor_set(x_24, 0, x_1); -return x_23; +lean_object* x_39; +x_39 = lean_ctor_get(x_29, 0); +lean_ctor_set(x_1, 0, x_39); +lean_ctor_set(x_29, 0, x_1); +return x_28; } else { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_24, 0); -lean_inc(x_35); -lean_dec(x_24); -lean_ctor_set(x_1, 0, x_35); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_1); -lean_ctor_set(x_23, 0, x_36); -return x_23; +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_29, 0); +lean_inc(x_40); +lean_dec(x_29); +lean_ctor_set(x_1, 0, x_40); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_1); +lean_ctor_set(x_28, 0, x_41); +return x_28; } } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = lean_ctor_get(x_23, 1); -lean_inc(x_37); -lean_dec(x_23); -x_38 = lean_ctor_get(x_24, 0); -lean_inc(x_38); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - x_39 = x_24; -} else { - lean_dec_ref(x_24); - x_39 = lean_box(0); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_42 = lean_ctor_get(x_28, 1); +lean_inc(x_42); +lean_dec(x_28); +x_43 = lean_ctor_get(x_29, 0); +lean_inc(x_43); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + x_44 = x_29; +} else { + lean_dec_ref(x_29); + x_44 = lean_box(0); } -lean_ctor_set(x_1, 0, x_38); -if (lean_is_scalar(x_39)) { - x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1, 0, x_43); +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(1, 1, 0); } else { - x_40 = x_39; + x_45 = x_44; } -lean_ctor_set(x_40, 0, x_1); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_37); -return x_41; +lean_ctor_set(x_45, 0, x_1); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_42); +return x_46; } } } else { -uint8_t x_42; +uint8_t x_47; lean_free_object(x_1); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -2676,191 +3754,216 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -x_42 = !lean_is_exclusive(x_23); -if (x_42 == 0) +x_47 = !lean_is_exclusive(x_28); +if (x_47 == 0) { -return x_23; +return x_28; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_23, 0); -x_44 = lean_ctor_get(x_23, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_23); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; +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; } } } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_46 = lean_ctor_get(x_1, 0); -x_47 = lean_ctor_get(x_1, 1); -x_48 = lean_ctor_get(x_1, 2); -x_49 = lean_ctor_get(x_1, 3); -x_50 = lean_ctor_get(x_1, 4); -x_51 = lean_ctor_get(x_1, 5); -x_52 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_53 = lean_ctor_get(x_1, 6); -x_54 = lean_ctor_get(x_1, 7); -x_55 = lean_ctor_get(x_1, 8); -x_56 = lean_ctor_get(x_1, 9); -x_57 = lean_ctor_get(x_1, 10); -x_58 = lean_ctor_get(x_1, 11); -x_59 = lean_ctor_get(x_1, 12); -x_60 = lean_ctor_get(x_1, 13); +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; 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_51 = lean_ctor_get(x_1, 0); +x_52 = lean_ctor_get(x_1, 1); +x_53 = lean_ctor_get(x_1, 2); +x_54 = lean_ctor_get(x_1, 3); +x_55 = lean_ctor_get(x_1, 4); +x_56 = lean_ctor_get(x_1, 5); +x_57 = lean_ctor_get_uint8(x_1, sizeof(void*)*19); +x_58 = lean_ctor_get(x_1, 6); +x_59 = lean_ctor_get(x_1, 7); +x_60 = lean_ctor_get(x_1, 8); +x_61 = lean_ctor_get(x_1, 9); +x_62 = lean_ctor_get(x_1, 10); +x_63 = lean_ctor_get(x_1, 11); +x_64 = lean_ctor_get(x_1, 12); +x_65 = lean_ctor_get(x_1, 13); +x_66 = lean_ctor_get(x_1, 14); +x_67 = lean_ctor_get(x_1, 15); +x_68 = lean_ctor_get(x_1, 16); +x_69 = lean_ctor_get(x_1, 17); +x_70 = lean_ctor_get(x_1, 18); +lean_inc(x_70); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); lean_inc(x_60); lean_inc(x_59); lean_inc(x_58); -lean_inc(x_57); lean_inc(x_56); lean_inc(x_55); lean_inc(x_54); lean_inc(x_53); +lean_inc(x_52); lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); lean_dec(x_1); -x_61 = l_Lean_Meta_Grind_injection_x3f(x_46, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_61) == 0) +x_71 = l_Lean_Meta_Grind_injection_x3f(x_51, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_71) == 0) { -lean_object* x_62; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -if (lean_obj_tag(x_62) == 0) +lean_object* x_72; +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +if (lean_obj_tag(x_72) == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +lean_dec(x_64); +lean_dec(x_63); +lean_dec(x_62); +lean_dec(x_61); lean_dec(x_60); lean_dec(x_59); lean_dec(x_58); -lean_dec(x_57); lean_dec(x_56); lean_dec(x_55); lean_dec(x_54); lean_dec(x_53); -lean_dec(x_51); -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_47); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_64 = x_61; +lean_dec(x_52); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_74 = x_71; } else { - lean_dec_ref(x_61); - x_64 = lean_box(0); + lean_dec_ref(x_71); + x_74 = lean_box(0); } -x_65 = lean_box(0); -if (lean_is_scalar(x_64)) { - x_66 = lean_alloc_ctor(0, 2, 0); +x_75 = lean_box(0); +if (lean_is_scalar(x_74)) { + x_76 = lean_alloc_ctor(0, 2, 0); } else { - x_66 = x_64; + x_76 = x_74; } -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_63); -return x_66; +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_73); +return x_76; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_67 = lean_ctor_get(x_61, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_68 = x_61; +lean_object* x_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_77 = lean_ctor_get(x_71, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_78 = x_71; } else { - lean_dec_ref(x_61); - x_68 = lean_box(0); + lean_dec_ref(x_71); + x_78 = lean_box(0); } -x_69 = lean_ctor_get(x_62, 0); -lean_inc(x_69); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - x_70 = x_62; +x_79 = lean_ctor_get(x_72, 0); +lean_inc(x_79); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + x_80 = x_72; } else { - lean_dec_ref(x_62); - x_70 = lean_box(0); -} -x_71 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_47); -lean_ctor_set(x_71, 2, x_48); -lean_ctor_set(x_71, 3, x_49); -lean_ctor_set(x_71, 4, x_50); -lean_ctor_set(x_71, 5, x_51); -lean_ctor_set(x_71, 6, x_53); -lean_ctor_set(x_71, 7, x_54); -lean_ctor_set(x_71, 8, x_55); -lean_ctor_set(x_71, 9, x_56); -lean_ctor_set(x_71, 10, x_57); -lean_ctor_set(x_71, 11, x_58); -lean_ctor_set(x_71, 12, x_59); -lean_ctor_set(x_71, 13, x_60); -lean_ctor_set_uint8(x_71, sizeof(void*)*14, x_52); -if (lean_is_scalar(x_70)) { - x_72 = lean_alloc_ctor(1, 1, 0); + lean_dec_ref(x_72); + x_80 = lean_box(0); +} +x_81 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_52); +lean_ctor_set(x_81, 2, x_53); +lean_ctor_set(x_81, 3, x_54); +lean_ctor_set(x_81, 4, x_55); +lean_ctor_set(x_81, 5, x_56); +lean_ctor_set(x_81, 6, x_58); +lean_ctor_set(x_81, 7, x_59); +lean_ctor_set(x_81, 8, x_60); +lean_ctor_set(x_81, 9, x_61); +lean_ctor_set(x_81, 10, x_62); +lean_ctor_set(x_81, 11, x_63); +lean_ctor_set(x_81, 12, x_64); +lean_ctor_set(x_81, 13, x_65); +lean_ctor_set(x_81, 14, x_66); +lean_ctor_set(x_81, 15, x_67); +lean_ctor_set(x_81, 16, x_68); +lean_ctor_set(x_81, 17, x_69); +lean_ctor_set(x_81, 18, x_70); +lean_ctor_set_uint8(x_81, sizeof(void*)*19, x_57); +if (lean_is_scalar(x_80)) { + x_82 = lean_alloc_ctor(1, 1, 0); } else { - x_72 = x_70; + x_82 = x_80; } -lean_ctor_set(x_72, 0, x_71); -if (lean_is_scalar(x_68)) { - x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +if (lean_is_scalar(x_78)) { + x_83 = lean_alloc_ctor(0, 2, 0); } else { - x_73 = x_68; + x_83 = x_78; } -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_67); -return x_73; +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_77); +return x_83; } } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_70); +lean_dec(x_69); +lean_dec(x_68); +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +lean_dec(x_64); +lean_dec(x_63); +lean_dec(x_62); +lean_dec(x_61); lean_dec(x_60); lean_dec(x_59); lean_dec(x_58); -lean_dec(x_57); lean_dec(x_56); lean_dec(x_55); lean_dec(x_54); lean_dec(x_53); -lean_dec(x_51); -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_47); -x_74 = lean_ctor_get(x_61, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_61, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_76 = x_61; +lean_dec(x_52); +x_84 = lean_ctor_get(x_71, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_71, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_86 = x_71; } else { - lean_dec_ref(x_61); - x_76 = lean_box(0); + lean_dec_ref(x_71); + x_86 = lean_box(0); } -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(1, 2, 0); } else { - x_77 = x_76; + x_87 = x_86; } -lean_ctor_set(x_77, 0, x_74); -lean_ctor_set(x_77, 1, x_75); -return x_77; +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_85); +return x_87; } } } @@ -3027,32 +4130,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_mk_ref(x_1, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; @@ -3156,85 +4234,57 @@ return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_1); -if (x_10 == 0) -{ -lean_object* x_11; -x_11 = lean_ctor_get(x_1, 1); -lean_dec(x_11); -lean_ctor_set(x_1, 1, x_9); -return x_1; -} -else -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_9); -return x_13; -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_intros_go___lambda__4___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__3___boxed), 9, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21, lean_object* x_22, lean_object* x_23, lean_object* x_24, lean_object* x_25, lean_object* x_26, lean_object* x_27, lean_object* x_28, lean_object* x_29, lean_object* x_30, lean_object* x_31, lean_object* x_32) { _start: { -lean_object* x_28; +lean_object* x_33; +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); +lean_inc(x_2); lean_inc(x_1); -x_28 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(x_1, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27); -if (lean_obj_tag(x_28) == 0) +x_33 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext(x_1, x_2, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_29; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -switch (lean_obj_tag(x_29)) { +lean_object* x_34; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +switch (lean_obj_tag(x_34)) { case 0: { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_28, 1); +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +lean_inc(x_31); lean_inc(x_30); -lean_dec(x_28); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -x_31 = l_Lean_MVarId_byContra_x3f(x_2, x_23, x_24, x_25, x_26, x_30); -if (lean_obj_tag(x_31) == 0) +lean_inc(x_29); +lean_inc(x_28); +x_36 = l_Lean_MVarId_byContra_x3f(x_3, x_28, x_29, x_30, x_31, x_35); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_32; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -if (lean_obj_tag(x_32) == 0) +lean_object* x_37; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -3243,86 +4293,96 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_st_ref_take(x_19, x_33); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_array_push(x_35, x_1); -x_38 = lean_st_ref_set(x_19, x_37, x_36); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_38, 0); -lean_dec(x_40); -x_41 = lean_box(0); -lean_ctor_set(x_38, 0, x_41); -return x_38; +lean_dec(x_2); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_st_ref_take(x_24, x_38); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_array_push(x_40, x_1); +x_43 = lean_st_ref_set(x_24, x_42, x_41); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 0); +lean_dec(x_45); +x_46 = lean_box(0); +lean_ctor_set(x_43, 0, x_46); +return x_43; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_38, 1); -lean_inc(x_42); -lean_dec(x_38); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -return x_44; +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_43, 1); +lean_inc(x_47); +lean_dec(x_43); +x_48 = lean_box(0); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +return x_49; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_dec(x_1); -x_45 = lean_ctor_get(x_31, 1); -lean_inc(x_45); +x_50 = lean_ctor_get(x_36, 1); +lean_inc(x_50); +lean_dec(x_36); +x_51 = lean_ctor_get(x_37, 0); +lean_inc(x_51); +lean_dec(x_37); +x_52 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_4); +lean_ctor_set(x_52, 2, x_5); +lean_ctor_set(x_52, 3, x_6); +lean_ctor_set(x_52, 4, x_7); +lean_ctor_set(x_52, 5, x_8); +lean_ctor_set(x_52, 6, x_10); +lean_ctor_set(x_52, 7, x_11); +lean_ctor_set(x_52, 8, x_12); +lean_ctor_set(x_52, 9, x_13); +lean_ctor_set(x_52, 10, x_14); +lean_ctor_set(x_52, 11, x_15); +lean_ctor_set(x_52, 12, x_16); +lean_ctor_set(x_52, 13, x_17); +lean_ctor_set(x_52, 14, x_18); +lean_ctor_set(x_52, 15, x_19); +lean_ctor_set(x_52, 16, x_20); +lean_ctor_set(x_52, 17, x_21); +lean_ctor_set(x_52, 18, x_22); +lean_ctor_set_uint8(x_52, sizeof(void*)*19, x_9); +x_53 = l_Lean_Meta_Grind_intros_go(x_2, x_52, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_50); +return x_53; +} +} +else +{ +uint8_t x_54; lean_dec(x_31); -x_46 = lean_ctor_get(x_32, 0); -lean_inc(x_46); -lean_dec(x_32); -x_47 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_3); -lean_ctor_set(x_47, 2, x_4); -lean_ctor_set(x_47, 3, x_5); -lean_ctor_set(x_47, 4, x_6); -lean_ctor_set(x_47, 5, x_7); -lean_ctor_set(x_47, 6, x_9); -lean_ctor_set(x_47, 7, x_10); -lean_ctor_set(x_47, 8, x_11); -lean_ctor_set(x_47, 9, x_12); -lean_ctor_set(x_47, 10, x_13); -lean_ctor_set(x_47, 11, x_14); -lean_ctor_set(x_47, 12, x_15); -lean_ctor_set(x_47, 13, x_16); -lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_8); -x_48 = l_Lean_Meta_Grind_intros_go(x_17, x_47, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_45); -return x_48; -} -} -else -{ -uint8_t x_49; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -3331,36 +4391,42 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_49 = !lean_is_exclusive(x_31); -if (x_49 == 0) +x_54 = !lean_is_exclusive(x_36); +if (x_54 == 0) { -return x_31; +return x_36; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_31, 0); -x_51 = lean_ctor_get(x_31, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_31); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_36, 0); +x_56 = lean_ctor_get(x_36, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_36); +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; } } } case 1: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -3368,222 +4434,227 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_53 = lean_ctor_get(x_28, 1); -lean_inc(x_53); -lean_dec(x_28); -x_54 = lean_ctor_get(x_29, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_29, 1); -lean_inc(x_55); -lean_dec(x_29); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_54); -lean_inc(x_55); -x_56 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_55, x_54, x_23, x_24, x_25, x_26, x_53); -if (lean_obj_tag(x_56) == 0) -{ -lean_object* x_57; -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -if (lean_obj_tag(x_57) == 0) -{ -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_56, 1); +x_58 = lean_ctor_get(x_33, 1); lean_inc(x_58); -lean_dec(x_56); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_54); -lean_inc(x_55); -x_59 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(x_55, x_54, x_23, x_24, x_25, x_26, x_58); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; -x_60 = lean_ctor_get(x_59, 0); +lean_dec(x_33); +x_59 = lean_ctor_get(x_34, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_34, 1); +lean_inc(x_60); +lean_dec(x_34); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_59); lean_inc(x_60); -if (lean_obj_tag(x_60) == 0) +x_61 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_60, x_59, x_28, x_29, x_30, x_31, x_58); +if (lean_obj_tag(x_61) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_ctor_get(x_55, 0); +lean_object* x_62; +x_62 = lean_ctor_get(x_61, 0); lean_inc(x_62); -x_63 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__1___boxed), 9, 1); -lean_closure_set(x_63, 0, x_55); -lean_inc(x_17); -x_64 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__2), 11, 2); -lean_closure_set(x_64, 0, x_54); -lean_closure_set(x_64, 1, x_17); -x_65 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_65, 0, x_63); -lean_closure_set(x_65, 1, x_64); -x_66 = l_Lean_Meta_Grind_intros_go___lambda__4___closed__1; -x_67 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_67, 0, x_65); -lean_closure_set(x_67, 1, x_66); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_59); +lean_inc(x_60); +x_64 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyInjection_x3f(x_60, x_59, x_28, x_29, x_30, x_31, x_63); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = lean_ctor_get(x_60, 0); +lean_inc(x_67); +x_68 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_68, 0, x_60); +lean_inc(x_2); +x_69 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_intros_go___lambda__1), 11, 2); +lean_closure_set(x_69, 0, x_59); +lean_closure_set(x_69, 1, x_2); +x_70 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_70, 0, x_68); +lean_closure_set(x_70, 1, x_69); +x_71 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_72 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_72, 0, x_70); +lean_closure_set(x_72, 1, x_71); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -x_68 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_62, x_67, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_61); -if (lean_obj_tag(x_68) == 0) +x_73 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_67, x_72, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_66); +if (lean_obj_tag(x_73) == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Lean_Meta_Grind_intros_go(x_17, x_69, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_70); -return x_71; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_Grind_intros_go(x_2, x_74, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_75); +return x_76; } else { -uint8_t x_72; +uint8_t x_77; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); -x_72 = !lean_is_exclusive(x_68); -if (x_72 == 0) +lean_dec(x_2); +x_77 = !lean_is_exclusive(x_73); +if (x_77 == 0) { -return x_68; +return x_73; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_68, 0); -x_74 = lean_ctor_get(x_68, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_68); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_73, 0); +x_79 = lean_ctor_get(x_73, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_73); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +return x_80; } } } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_dec(x_55); -lean_dec(x_54); -x_76 = lean_ctor_get(x_59, 1); -lean_inc(x_76); -lean_dec(x_59); -x_77 = lean_ctor_get(x_60, 0); -lean_inc(x_77); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_dec(x_60); -x_78 = l_Lean_Meta_Grind_intros_go(x_17, x_77, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_76); -return x_78; +lean_dec(x_59); +x_81 = lean_ctor_get(x_64, 1); +lean_inc(x_81); +lean_dec(x_64); +x_82 = lean_ctor_get(x_65, 0); +lean_inc(x_82); +lean_dec(x_65); +x_83 = l_Lean_Meta_Grind_intros_go(x_2, x_82, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_81); +return x_83; } } else { -uint8_t x_79; -lean_dec(x_55); -lean_dec(x_54); +uint8_t x_84; +lean_dec(x_60); +lean_dec(x_59); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); -x_79 = !lean_is_exclusive(x_59); -if (x_79 == 0) +lean_dec(x_2); +x_84 = !lean_is_exclusive(x_64); +if (x_84 == 0) { -return x_59; +return x_64; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_59, 0); -x_81 = lean_ctor_get(x_59, 1); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_59); -x_82 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -return x_82; +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_64, 0); +x_86 = lean_ctor_get(x_64, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_64); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; } } } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_55); -lean_dec(x_54); -x_83 = lean_ctor_get(x_56, 1); -lean_inc(x_83); -lean_dec(x_56); -x_84 = lean_ctor_get(x_57, 0); -lean_inc(x_84); -lean_dec(x_57); -x_85 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_17, x_84, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_83); -return x_85; +lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_dec(x_60); +lean_dec(x_59); +x_88 = lean_ctor_get(x_61, 1); +lean_inc(x_88); +lean_dec(x_61); +x_89 = lean_ctor_get(x_62, 0); +lean_inc(x_89); +lean_dec(x_62); +x_90 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_2, x_89, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_88); +return x_90; } } else { -uint8_t x_86; -lean_dec(x_55); -lean_dec(x_54); +uint8_t x_91; +lean_dec(x_60); +lean_dec(x_59); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); -x_86 = !lean_is_exclusive(x_56); -if (x_86 == 0) +lean_dec(x_2); +x_91 = !lean_is_exclusive(x_61); +if (x_91 == 0) { -return x_56; +return x_61; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_56, 0); -x_88 = lean_ctor_get(x_56, 1); -lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_56); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_61, 0); +x_93 = lean_ctor_get(x_61, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_61); +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; } } } case 2: { -lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -3591,26 +4662,31 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_90 = lean_ctor_get(x_28, 1); -lean_inc(x_90); -lean_dec(x_28); -x_91 = lean_ctor_get(x_29, 0); -lean_inc(x_91); -lean_dec(x_29); -x_92 = l_Lean_Meta_Grind_intros_go(x_17, x_91, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_90); -return x_92; +x_95 = lean_ctor_get(x_33, 1); +lean_inc(x_95); +lean_dec(x_33); +x_96 = lean_ctor_get(x_34, 0); +lean_inc(x_96); +lean_dec(x_34); +x_97 = l_Lean_Meta_Grind_intros_go(x_2, x_96, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_95); +return x_97; } default: { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); @@ -3618,85 +4694,84 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_93 = lean_ctor_get(x_28, 1); -lean_inc(x_93); -lean_dec(x_28); -x_94 = lean_ctor_get(x_29, 0); -lean_inc(x_94); -x_95 = lean_ctor_get(x_29, 1); -lean_inc(x_95); -lean_dec(x_29); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_95); -x_96 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_95, x_94, x_23, x_24, x_25, x_26, x_93); -if (lean_obj_tag(x_96) == 0) +x_98 = lean_ctor_get(x_33, 1); +lean_inc(x_98); +lean_dec(x_33); +x_99 = lean_ctor_get(x_34, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_34, 1); +lean_inc(x_100); +lean_dec(x_34); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_100); +x_101 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_applyCases_x3f(x_100, x_99, x_28, x_29, x_30, x_31, x_98); +if (lean_obj_tag(x_101) == 0) { -lean_object* x_97; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -if (lean_obj_tag(x_97) == 0) +lean_object* x_102; +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); +if (lean_obj_tag(x_102) == 0) { -lean_object* x_98; lean_object* x_99; -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = l_Lean_Meta_Grind_intros_go(x_17, x_95, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_98); -return x_99; +lean_object* x_103; lean_object* x_104; +x_103 = lean_ctor_get(x_101, 1); +lean_inc(x_103); +lean_dec(x_101); +x_104 = l_Lean_Meta_Grind_intros_go(x_2, x_100, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_103); +return x_104; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -lean_dec(x_95); -x_100 = lean_ctor_get(x_96, 1); -lean_inc(x_100); -lean_dec(x_96); -x_101 = lean_ctor_get(x_97, 0); -lean_inc(x_101); -lean_dec(x_97); -x_102 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_17, x_101, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_100); -return x_102; +lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_100); +x_105 = lean_ctor_get(x_101, 1); +lean_inc(x_105); +lean_dec(x_101); +x_106 = lean_ctor_get(x_102, 0); +lean_inc(x_106); +lean_dec(x_102); +x_107 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_2, x_106, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_105); +return x_107; } } else { -uint8_t x_103; -lean_dec(x_95); +uint8_t x_108; +lean_dec(x_100); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_17); -x_103 = !lean_is_exclusive(x_96); -if (x_103 == 0) +lean_dec(x_2); +x_108 = !lean_is_exclusive(x_101); +if (x_108 == 0) { -return x_96; +return x_101; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_96, 0); -x_105 = lean_ctor_get(x_96, 1); -lean_inc(x_105); -lean_inc(x_104); -lean_dec(x_96); -x_106 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_106, 0, x_104); -lean_ctor_set(x_106, 1, x_105); -return x_106; +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_101, 0); +x_110 = lean_ctor_get(x_101, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_101); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } } @@ -3704,14 +4779,19 @@ return x_106; } else { -uint8_t x_107; +uint8_t x_112; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -3720,78 +4800,437 @@ lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_112 = !lean_is_exclusive(x_33); +if (x_112 == 0) +{ +return x_33; +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_33, 0); +x_114 = lean_ctor_get(x_33, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_33); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +return x_115; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*19); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 3); +lean_inc(x_16); +x_17 = lean_ctor_get(x_2, 4); +lean_inc(x_17); +x_18 = lean_ctor_get(x_2, 5); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 6); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 7); +lean_inc(x_20); +x_21 = lean_ctor_get(x_2, 8); +lean_inc(x_21); +x_22 = lean_ctor_get(x_2, 9); +lean_inc(x_22); +x_23 = lean_ctor_get(x_2, 10); +lean_inc(x_23); +x_24 = lean_ctor_get(x_2, 11); +lean_inc(x_24); +x_25 = lean_ctor_get(x_2, 12); +lean_inc(x_25); +x_26 = lean_ctor_get(x_2, 13); +lean_inc(x_26); +x_27 = lean_ctor_get(x_2, 14); +lean_inc(x_27); +x_28 = lean_ctor_get(x_2, 15); +lean_inc(x_28); +x_29 = lean_ctor_get(x_2, 16); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 17); +lean_inc(x_30); +x_31 = lean_ctor_get(x_2, 18); +lean_inc(x_31); +x_32 = lean_box(0); +x_33 = l_Lean_Meta_Grind_intros_go___lambda__2(x_2, x_1, x_13, x_14, x_15, x_16, x_17, x_18, x_12, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_33; +} +else +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_11); +return x_35; +} +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; +lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; +lean_object* x_27 = _args[26]; +lean_object* x_28 = _args[27]; +lean_object* x_29 = _args[28]; +lean_object* x_30 = _args[29]; +lean_object* x_31 = _args[30]; +lean_object* x_32 = _args[31]; +_start: +{ +uint8_t x_33; lean_object* x_34; +x_33 = lean_unbox(x_9); +lean_dec(x_9); +x_34 = l_Lean_Meta_Grind_intros_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_33, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27, x_28, x_29, x_30, x_31, x_32); +lean_dec(x_24); +lean_dec(x_23); +return x_34; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_intros_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +static lean_object* _init_l_Lean_Meta_Grind_intros___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = l_Lean_Meta_Grind_intros___closed__1; +x_12 = lean_st_mk_ref(x_11, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Meta_Grind_intros_go(x_1, x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_get(x_13, x_16); +lean_dec(x_13); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_array_to_list(x_19); +lean_ctor_set(x_17, 0, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_17); +x_23 = lean_array_to_list(x_21); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +else +{ +uint8_t x_25; +lean_dec(x_13); +x_25 = !lean_is_exclusive(x_15); +if (x_25 == 0) +{ +return x_15; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_6); +x_13 = l_Lean_Meta_Grind_simp(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_Meta_Simp_Result_getProof(x_14, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_20 = l_Lean_Meta_mkEqMP(x_18, x_2, x_8, x_9, x_10, x_11, x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +lean_inc(x_4); +x_23 = l_Lean_Meta_Grind_add(x_16, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_st_ref_get(x_4, x_24); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_25, 1); +x_28 = lean_st_ref_get(x_4, x_27); +lean_dec(x_4); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +lean_object* x_30; +x_30 = lean_ctor_get(x_28, 0); +lean_ctor_set(x_25, 1, x_30); +lean_ctor_set(x_28, 0, x_25); +return x_28; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_28, 0); +x_32 = lean_ctor_get(x_28, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_28); +lean_ctor_set(x_25, 1, x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_25); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_34 = lean_ctor_get(x_25, 0); +x_35 = lean_ctor_get(x_25, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_25); +x_36 = lean_st_ref_get(x_4, x_35); +lean_dec(x_4); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_39 = x_36; +} else { + lean_dec_ref(x_36); + x_39 = lean_box(0); +} +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_37); +if (lean_is_scalar(x_39)) { + x_41 = lean_alloc_ctor(0, 2, 0); +} else { + x_41 = x_39; +} +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +return x_41; +} +} +else +{ +uint8_t x_42; +lean_dec(x_4); +x_42 = !lean_is_exclusive(x_23); +if (x_42 == 0) +{ +return x_23; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_23, 0); +x_44 = lean_ctor_get(x_23, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_23); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; +} +} +} +else +{ +uint8_t x_46; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_107 = !lean_is_exclusive(x_28); -if (x_107 == 0) +x_46 = !lean_is_exclusive(x_20); +if (x_46 == 0) { -return x_28; +return x_20; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_28, 0); -x_109 = lean_ctor_get(x_28, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_28); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} -} +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_20, 0); +x_48 = lean_ctor_get(x_20, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_20); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = lean_ctor_get_uint8(x_2, sizeof(void*)*14); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_13 = lean_ctor_get(x_2, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_2, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_2, 2); -lean_inc(x_15); -x_16 = lean_ctor_get(x_2, 3); -lean_inc(x_16); -x_17 = lean_ctor_get(x_2, 4); -lean_inc(x_17); -x_18 = lean_ctor_get(x_2, 5); -lean_inc(x_18); -x_19 = lean_ctor_get(x_2, 6); -lean_inc(x_19); -x_20 = lean_ctor_get(x_2, 7); -lean_inc(x_20); -x_21 = lean_ctor_get(x_2, 8); -lean_inc(x_21); -x_22 = lean_ctor_get(x_2, 9); -lean_inc(x_22); -x_23 = lean_ctor_get(x_2, 10); -lean_inc(x_23); -x_24 = lean_ctor_get(x_2, 11); -lean_inc(x_24); -x_25 = lean_ctor_get(x_2, 12); -lean_inc(x_25); -x_26 = lean_ctor_get(x_2, 13); -lean_inc(x_26); -x_27 = lean_box(0); -x_28 = l_Lean_Meta_Grind_intros_go___lambda__4(x_2, x_13, x_14, x_15, x_16, x_17, x_18, x_12, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_1, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_28; } else { -lean_object* x_29; lean_object* x_30; +uint8_t x_50; +lean_dec(x_16); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -3799,39 +5238,34 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_11); -return x_30; -} -} -} -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +x_50 = !lean_is_exclusive(x_17); +if (x_50 == 0) { -lean_object* x_12; -x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -return x_12; -} +return x_17; } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_List_forM___at_Lean_Meta_Grind_intros_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -return x_12; +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_17, 0); +x_52 = lean_ctor_get(x_17, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_17); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +} +else { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_intros_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_54; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -3839,231 +5273,351 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_10; -} +x_54 = !lean_is_exclusive(x_13); +if (x_54 == 0) +{ +return x_13; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_intros_go___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_13, 0); +x_56 = lean_ctor_get(x_13, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_13); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___lambda__4___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -lean_object* x_20 = _args[19]; -lean_object* x_21 = _args[20]; -lean_object* x_22 = _args[21]; -lean_object* x_23 = _args[22]; -lean_object* x_24 = _args[23]; -lean_object* x_25 = _args[24]; -lean_object* x_26 = _args[25]; -lean_object* x_27 = _args[26]; -_start: -{ -uint8_t x_28; lean_object* x_29; -x_28 = lean_unbox(x_8); -lean_dec(x_8); -x_29 = l_Lean_Meta_Grind_intros_go___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26, x_27); -lean_dec(x_19); -lean_dec(x_18); -return x_29; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l_Lean_Meta_Grind_assertAt___closed__1() { _start: { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_intros_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_3); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("h", 1, 1); +return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_intros___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_assertAt___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; +x_2 = l_Lean_Meta_Grind_assertAt___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_intros(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = l_Lean_Meta_Grind_intros___closed__1; -x_12 = lean_st_mk_ref(x_11, x_10); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_isCasesCandidate(x_2, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Meta_Grind_intros_go(x_2, x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_15) == 0) +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) { -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_st_ref_get(x_13, x_16); -lean_dec(x_13); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_array_to_list(x_19); -lean_ctor_set(x_17, 0, x_20); -return x_17; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_17 = lean_ctor_get(x_13, 1); +x_18 = lean_ctor_get(x_13, 0); +lean_dec(x_18); +x_19 = lean_ctor_get(x_4, 0); +lean_inc(x_19); +x_20 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_20, 0, x_4); +x_21 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertAt___lambda__1), 12, 3); +lean_closure_set(x_21, 0, x_2); +lean_closure_set(x_21, 1, x_1); +lean_closure_set(x_21, 2, x_3); +x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_21); +x_23 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_24, 0, x_22); +lean_closure_set(x_24, 1, x_23); +x_25 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_19, x_24, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get_uint8(x_26, sizeof(void*)*19); +if (x_27 == 0) +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_25); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_25, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 1, x_30); +lean_ctor_set(x_13, 0, x_26); +lean_ctor_set(x_25, 0, x_13); +return x_25; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_17, 0); -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_17); -x_23 = lean_array_to_list(x_21); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_32 = lean_box(0); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 1, x_32); +lean_ctor_set(x_13, 0, x_26); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_13); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +else +{ +uint8_t x_34; +lean_dec(x_26); +lean_free_object(x_13); +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_25, 0); +lean_dec(x_35); +x_36 = lean_box(0); +lean_ctor_set(x_25, 0, x_36); +return x_25; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_25, 1); +lean_inc(x_37); +lean_dec(x_25); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +return x_39; +} +} +} +else +{ +uint8_t x_40; +lean_free_object(x_13); +x_40 = !lean_is_exclusive(x_25); +if (x_40 == 0) +{ +return x_25; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_25, 0); +x_42 = lean_ctor_get(x_25, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_25); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} } } else { -uint8_t x_25; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_44 = lean_ctor_get(x_13, 1); +lean_inc(x_44); lean_dec(x_13); -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) +x_45 = lean_ctor_get(x_4, 0); +lean_inc(x_45); +x_46 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__2___boxed), 9, 1); +lean_closure_set(x_46, 0, x_4); +x_47 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertAt___lambda__1), 12, 3); +lean_closure_set(x_47, 0, x_2); +lean_closure_set(x_47, 1, x_1); +lean_closure_set(x_47, 2, x_3); +x_48 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_48, 0, x_46); +lean_closure_set(x_48, 1, x_47); +x_49 = l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1; +x_50 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_50, 0, x_48); +lean_closure_set(x_50, 1, x_49); +x_51 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_45, x_50, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); +if (lean_obj_tag(x_51) == 0) { -return x_15; +lean_object* x_52; uint8_t x_53; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get_uint8(x_52, sizeof(void*)*19); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_55 = x_51; +} else { + lean_dec_ref(x_51); + x_55 = lean_box(0); +} +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_52); +lean_ctor_set(x_57, 1, x_56); +if (lean_is_scalar(x_55)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_55; +} +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_54); +return x_58; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_52); +x_59 = lean_ctor_get(x_51, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_60 = x_51; +} else { + lean_dec_ref(x_51); + x_60 = lean_box(0); } +x_61 = lean_box(0); +if (lean_is_scalar(x_60)) { + x_62 = lean_alloc_ctor(0, 2, 0); +} else { + x_62 = x_60; } +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_59); +return x_62; } } -static lean_object* _init_l_Lean_Meta_Grind_assertAt___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("h", 1, 1); -return x_1; +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_51, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_51, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_65 = x_51; +} else { + lean_dec_ref(x_51); + x_65 = lean_box(0); } +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(1, 2, 0); +} else { + x_66 = x_65; } -static lean_object* _init_l_Lean_Meta_Grind_assertAt___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_assertAt___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_ctor_set(x_66, 0, x_63); +lean_ctor_set(x_66, 1, x_64); +return x_66; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +} +else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_13 = l_Lean_Meta_Grind_assertAt___closed__2; -x_14 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_13, x_10, x_11, x_12); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_1); -if (x_17 == 0) +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_67 = lean_ctor_get(x_13, 1); +lean_inc(x_67); +lean_dec(x_13); +x_68 = l_Lean_Meta_Grind_assertAt___closed__2; +x_69 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_68, x_10, x_11, x_67); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = !lean_is_exclusive(x_4); +if (x_72 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_18 = lean_ctor_get(x_1, 0); -x_19 = lean_ctor_get(x_1, 1); -x_20 = lean_ctor_get(x_1, 2); -x_21 = lean_ctor_get(x_1, 3); -x_22 = lean_ctor_get(x_1, 4); -x_23 = lean_ctor_get(x_1, 5); -x_24 = lean_ctor_get(x_1, 6); -x_25 = lean_ctor_get(x_1, 7); -x_26 = lean_ctor_get(x_1, 8); -x_27 = lean_ctor_get(x_1, 9); -x_28 = lean_ctor_get(x_1, 10); -x_29 = lean_ctor_get(x_1, 11); -x_30 = lean_ctor_get(x_1, 12); -x_31 = lean_ctor_get(x_1, 13); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_73 = lean_ctor_get(x_4, 0); +x_74 = lean_ctor_get(x_4, 1); +x_75 = lean_ctor_get(x_4, 2); +x_76 = lean_ctor_get(x_4, 3); +x_77 = lean_ctor_get(x_4, 4); +x_78 = lean_ctor_get(x_4, 5); +x_79 = lean_ctor_get(x_4, 6); +x_80 = lean_ctor_get(x_4, 7); +x_81 = lean_ctor_get(x_4, 8); +x_82 = lean_ctor_get(x_4, 9); +x_83 = lean_ctor_get(x_4, 10); +x_84 = lean_ctor_get(x_4, 11); +x_85 = lean_ctor_get(x_4, 12); +x_86 = lean_ctor_get(x_4, 13); +x_87 = lean_ctor_get(x_4, 14); +x_88 = lean_ctor_get(x_4, 15); +x_89 = lean_ctor_get(x_4, 16); +x_90 = lean_ctor_get(x_4, 17); +x_91 = lean_ctor_get(x_4, 18); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_32 = l_Lean_MVarId_assert(x_18, x_15, x_3, x_2, x_8, x_9, x_10, x_11, x_16); -if (lean_obj_tag(x_32) == 0) +x_92 = l_Lean_MVarId_assert(x_73, x_70, x_2, x_1, x_8, x_9, x_10, x_11, x_71); +if (lean_obj_tag(x_92) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -lean_ctor_set(x_1, 0, x_33); -x_35 = l_Lean_Meta_Grind_intros(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); -return x_35; +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +lean_ctor_set(x_4, 0, x_93); +x_95 = l_Lean_Meta_Grind_intros(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_94); +return x_95; } else { -uint8_t x_36; -lean_free_object(x_1); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_28); -lean_dec(x_27); -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); +uint8_t x_96; +lean_free_object(x_4); +lean_dec(x_91); +lean_dec(x_90); +lean_dec(x_89); +lean_dec(x_88); +lean_dec(x_87); +lean_dec(x_86); +lean_dec(x_85); +lean_dec(x_84); +lean_dec(x_83); +lean_dec(x_82); +lean_dec(x_81); +lean_dec(x_80); +lean_dec(x_79); +lean_dec(x_78); +lean_dec(x_77); +lean_dec(x_76); +lean_dec(x_75); +lean_dec(x_74); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4071,108 +5625,128 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_36 = !lean_is_exclusive(x_32); -if (x_36 == 0) +lean_dec(x_3); +x_96 = !lean_is_exclusive(x_92); +if (x_96 == 0) { -return x_32; +return x_92; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_32, 0); -x_38 = lean_ctor_get(x_32, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_32); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_92, 0); +x_98 = lean_ctor_get(x_92, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_92); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_40 = lean_ctor_get(x_1, 0); -x_41 = lean_ctor_get(x_1, 1); -x_42 = lean_ctor_get(x_1, 2); -x_43 = lean_ctor_get(x_1, 3); -x_44 = lean_ctor_get(x_1, 4); -x_45 = lean_ctor_get(x_1, 5); -x_46 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_47 = lean_ctor_get(x_1, 6); -x_48 = lean_ctor_get(x_1, 7); -x_49 = lean_ctor_get(x_1, 8); -x_50 = lean_ctor_get(x_1, 9); -x_51 = lean_ctor_get(x_1, 10); -x_52 = lean_ctor_get(x_1, 11); -x_53 = lean_ctor_get(x_1, 12); -x_54 = lean_ctor_get(x_1, 13); -lean_inc(x_54); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_1); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_100 = lean_ctor_get(x_4, 0); +x_101 = lean_ctor_get(x_4, 1); +x_102 = lean_ctor_get(x_4, 2); +x_103 = lean_ctor_get(x_4, 3); +x_104 = lean_ctor_get(x_4, 4); +x_105 = lean_ctor_get(x_4, 5); +x_106 = lean_ctor_get_uint8(x_4, sizeof(void*)*19); +x_107 = lean_ctor_get(x_4, 6); +x_108 = lean_ctor_get(x_4, 7); +x_109 = lean_ctor_get(x_4, 8); +x_110 = lean_ctor_get(x_4, 9); +x_111 = lean_ctor_get(x_4, 10); +x_112 = lean_ctor_get(x_4, 11); +x_113 = lean_ctor_get(x_4, 12); +x_114 = lean_ctor_get(x_4, 13); +x_115 = lean_ctor_get(x_4, 14); +x_116 = lean_ctor_get(x_4, 15); +x_117 = lean_ctor_get(x_4, 16); +x_118 = lean_ctor_get(x_4, 17); +x_119 = lean_ctor_get(x_4, 18); +lean_inc(x_119); +lean_inc(x_118); +lean_inc(x_117); +lean_inc(x_116); +lean_inc(x_115); +lean_inc(x_114); +lean_inc(x_113); +lean_inc(x_112); +lean_inc(x_111); +lean_inc(x_110); +lean_inc(x_109); +lean_inc(x_108); +lean_inc(x_107); +lean_inc(x_105); +lean_inc(x_104); +lean_inc(x_103); +lean_inc(x_102); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_4); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_55 = l_Lean_MVarId_assert(x_40, x_15, x_3, x_2, x_8, x_9, x_10, x_11, x_16); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_41); -lean_ctor_set(x_58, 2, x_42); -lean_ctor_set(x_58, 3, x_43); -lean_ctor_set(x_58, 4, x_44); -lean_ctor_set(x_58, 5, x_45); -lean_ctor_set(x_58, 6, x_47); -lean_ctor_set(x_58, 7, x_48); -lean_ctor_set(x_58, 8, x_49); -lean_ctor_set(x_58, 9, x_50); -lean_ctor_set(x_58, 10, x_51); -lean_ctor_set(x_58, 11, x_52); -lean_ctor_set(x_58, 12, x_53); -lean_ctor_set(x_58, 13, x_54); -lean_ctor_set_uint8(x_58, sizeof(void*)*14, x_46); -x_59 = l_Lean_Meta_Grind_intros(x_58, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_57); -return x_59; +x_120 = l_Lean_MVarId_assert(x_100, x_70, x_2, x_1, x_8, x_9, x_10, x_11, x_71); +if (lean_obj_tag(x_120) == 0) +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +lean_dec(x_120); +x_123 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_101); +lean_ctor_set(x_123, 2, x_102); +lean_ctor_set(x_123, 3, x_103); +lean_ctor_set(x_123, 4, x_104); +lean_ctor_set(x_123, 5, x_105); +lean_ctor_set(x_123, 6, x_107); +lean_ctor_set(x_123, 7, x_108); +lean_ctor_set(x_123, 8, x_109); +lean_ctor_set(x_123, 9, x_110); +lean_ctor_set(x_123, 10, x_111); +lean_ctor_set(x_123, 11, x_112); +lean_ctor_set(x_123, 12, x_113); +lean_ctor_set(x_123, 13, x_114); +lean_ctor_set(x_123, 14, x_115); +lean_ctor_set(x_123, 15, x_116); +lean_ctor_set(x_123, 16, x_117); +lean_ctor_set(x_123, 17, x_118); +lean_ctor_set(x_123, 18, x_119); +lean_ctor_set_uint8(x_123, sizeof(void*)*19, x_106); +x_124 = l_Lean_Meta_Grind_intros(x_3, x_123, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_122); +return x_124; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -lean_dec(x_54); -lean_dec(x_53); -lean_dec(x_52); -lean_dec(x_51); -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_47); -lean_dec(x_45); -lean_dec(x_44); -lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_41); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +lean_dec(x_119); +lean_dec(x_118); +lean_dec(x_117); +lean_dec(x_116); +lean_dec(x_115); +lean_dec(x_114); +lean_dec(x_113); +lean_dec(x_112); +lean_dec(x_111); +lean_dec(x_110); +lean_dec(x_109); +lean_dec(x_108); +lean_dec(x_107); +lean_dec(x_105); +lean_dec(x_104); +lean_dec(x_103); +lean_dec(x_102); +lean_dec(x_101); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4180,39 +5754,40 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_60 = lean_ctor_get(x_55, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_55, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_62 = x_55; +lean_dec(x_3); +x_125 = lean_ctor_get(x_120, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_120, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_127 = x_120; } else { - lean_dec_ref(x_55); - x_62 = lean_box(0); + lean_dec_ref(x_120); + x_127 = lean_box(0); } -if (lean_is_scalar(x_62)) { - x_63 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_127)) { + x_128 = lean_alloc_ctor(1, 2, 0); } else { - x_63 = x_62; + x_128 = x_127; } -lean_ctor_set(x_63, 0, x_60); -lean_ctor_set(x_63, 1, x_61); -return x_63; +lean_ctor_set(x_128, 0, x_125); +lean_ctor_set(x_128, 1, x_126); +return x_128; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertNext_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; x_10 = !lean_is_exclusive(x_1); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; x_11 = lean_ctor_get(x_1, 0); x_12 = lean_ctor_get(x_1, 1); x_13 = lean_ctor_get(x_1, 2); @@ -4227,11 +5802,21 @@ x_21 = lean_ctor_get(x_1, 10); x_22 = lean_ctor_get(x_1, 11); x_23 = lean_ctor_get(x_1, 12); x_24 = lean_ctor_get(x_1, 13); -x_25 = l_Std_Queue_dequeue_x3f___rarg(x_24); -if (lean_obj_tag(x_25) == 0) +x_25 = lean_ctor_get(x_1, 14); +x_26 = lean_ctor_get(x_1, 15); +x_27 = lean_ctor_get(x_1, 16); +x_28 = lean_ctor_get(x_1, 17); +x_29 = lean_ctor_get(x_1, 18); +x_30 = l_Std_Queue_dequeue_x3f___rarg(x_25); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_26; lean_object* x_27; +lean_object* x_31; lean_object* x_32; lean_free_object(x_1); +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); lean_dec(x_21); @@ -4252,208 +5837,223 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_26 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_9); -return x_27; +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_9); +return x_32; } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_25); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_29 = lean_ctor_get(x_25, 0); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -lean_ctor_set(x_1, 13, x_31); -x_32 = lean_ctor_get(x_30, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_30, 1); -lean_inc(x_33); -x_34 = lean_ctor_get(x_30, 2); -lean_inc(x_34); -lean_dec(x_30); -x_35 = l_Lean_Meta_Grind_assertAt(x_1, x_32, x_33, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_35) == 0) -{ -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +uint8_t x_33; +x_33 = !lean_is_exclusive(x_30); +if (x_33 == 0) { -lean_object* x_37; +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_34 = lean_ctor_get(x_30, 0); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); x_37 = lean_ctor_get(x_35, 0); -lean_ctor_set(x_25, 0, x_37); -lean_ctor_set(x_35, 0, x_25); -return x_35; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_35, 0); -x_39 = lean_ctor_get(x_35, 1); -lean_inc(x_39); +lean_inc(x_37); +x_38 = lean_ctor_get(x_35, 1); lean_inc(x_38); +x_39 = lean_ctor_get(x_35, 2); +lean_inc(x_39); lean_dec(x_35); -lean_ctor_set(x_25, 0, x_38); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_25); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} -} -else +lean_ctor_set(x_1, 14, x_36); +x_40 = l_Lean_Meta_Grind_assertAt(x_37, x_38, x_39, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_40) == 0) { uint8_t x_41; -lean_free_object(x_25); -x_41 = !lean_is_exclusive(x_35); +x_41 = !lean_is_exclusive(x_40); if (x_41 == 0) { -return x_35; +lean_object* x_42; +x_42 = lean_ctor_get(x_40, 0); +lean_ctor_set(x_30, 0, x_42); +lean_ctor_set(x_40, 0, x_30); +return x_40; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_35, 0); -x_43 = lean_ctor_get(x_35, 1); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_40, 0); +x_44 = lean_ctor_get(x_40, 1); +lean_inc(x_44); lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_35); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_dec(x_40); +lean_ctor_set(x_30, 0, x_43); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_30); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } +else +{ +uint8_t x_46; +lean_free_object(x_30); +x_46 = !lean_is_exclusive(x_40); +if (x_46 == 0) +{ +return x_40; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_45 = lean_ctor_get(x_25, 0); -lean_inc(x_45); -lean_dec(x_25); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -lean_ctor_set(x_1, 13, x_47); -x_48 = lean_ctor_get(x_46, 0); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_40, 0); +x_48 = lean_ctor_get(x_40, 1); lean_inc(x_48); -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -x_50 = lean_ctor_get(x_46, 2); -lean_inc(x_50); -lean_dec(x_46); -x_51 = l_Lean_Meta_Grind_assertAt(x_1, x_48, x_49, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_51) == 0) +lean_inc(x_47); +lean_dec(x_40); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +} +else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_52 = lean_ctor_get(x_51, 0); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_50 = lean_ctor_get(x_30, 0); +lean_inc(x_50); +lean_dec(x_30); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); +lean_dec(x_50); +x_53 = lean_ctor_get(x_51, 0); lean_inc(x_53); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_54 = x_51; +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +x_55 = lean_ctor_get(x_51, 2); +lean_inc(x_55); +lean_dec(x_51); +lean_ctor_set(x_1, 14, x_52); +x_56 = l_Lean_Meta_Grind_assertAt(x_53, x_54, x_55, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; } else { - lean_dec_ref(x_51); - x_54 = lean_box(0); + lean_dec_ref(x_56); + x_59 = lean_box(0); } -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_52); -if (lean_is_scalar(x_54)) { - x_56 = lean_alloc_ctor(0, 2, 0); +x_60 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_60, 0, x_57); +if (lean_is_scalar(x_59)) { + x_61 = lean_alloc_ctor(0, 2, 0); } else { - x_56 = x_54; + x_61 = x_59; } -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -return x_56; +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_58); +return x_61; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_57 = lean_ctor_get(x_51, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_51, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_51)) { - lean_ctor_release(x_51, 0); - lean_ctor_release(x_51, 1); - x_59 = x_51; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_56, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_56, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_64 = x_56; } else { - lean_dec_ref(x_51); - x_59 = lean_box(0); + lean_dec_ref(x_56); + x_64 = lean_box(0); } -if (lean_is_scalar(x_59)) { - x_60 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(1, 2, 0); } else { - x_60 = x_59; + x_65 = x_64; } -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_58); -return x_60; +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_63); +return x_65; } } } } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_61 = lean_ctor_get(x_1, 0); -x_62 = lean_ctor_get(x_1, 1); -x_63 = lean_ctor_get(x_1, 2); -x_64 = lean_ctor_get(x_1, 3); -x_65 = lean_ctor_get(x_1, 4); -x_66 = lean_ctor_get(x_1, 5); -x_67 = lean_ctor_get_uint8(x_1, sizeof(void*)*14); -x_68 = lean_ctor_get(x_1, 6); -x_69 = lean_ctor_get(x_1, 7); -x_70 = lean_ctor_get(x_1, 8); -x_71 = lean_ctor_get(x_1, 9); -x_72 = lean_ctor_get(x_1, 10); -x_73 = lean_ctor_get(x_1, 11); -x_74 = lean_ctor_get(x_1, 12); -x_75 = lean_ctor_get(x_1, 13); +lean_object* x_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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_66 = lean_ctor_get(x_1, 0); +x_67 = lean_ctor_get(x_1, 1); +x_68 = lean_ctor_get(x_1, 2); +x_69 = lean_ctor_get(x_1, 3); +x_70 = lean_ctor_get(x_1, 4); +x_71 = lean_ctor_get(x_1, 5); +x_72 = lean_ctor_get_uint8(x_1, sizeof(void*)*19); +x_73 = lean_ctor_get(x_1, 6); +x_74 = lean_ctor_get(x_1, 7); +x_75 = lean_ctor_get(x_1, 8); +x_76 = lean_ctor_get(x_1, 9); +x_77 = lean_ctor_get(x_1, 10); +x_78 = lean_ctor_get(x_1, 11); +x_79 = lean_ctor_get(x_1, 12); +x_80 = lean_ctor_get(x_1, 13); +x_81 = lean_ctor_get(x_1, 14); +x_82 = lean_ctor_get(x_1, 15); +x_83 = lean_ctor_get(x_1, 16); +x_84 = lean_ctor_get(x_1, 17); +x_85 = lean_ctor_get(x_1, 18); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_inc(x_79); +lean_inc(x_78); +lean_inc(x_77); +lean_inc(x_76); lean_inc(x_75); lean_inc(x_74); lean_inc(x_73); -lean_inc(x_72); lean_inc(x_71); lean_inc(x_70); lean_inc(x_69); lean_inc(x_68); +lean_inc(x_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_dec(x_1); -x_76 = l_Std_Queue_dequeue_x3f___rarg(x_75); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; +x_86 = l_Std_Queue_dequeue_x3f___rarg(x_81); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; +lean_dec(x_85); +lean_dec(x_84); +lean_dec(x_83); +lean_dec(x_82); +lean_dec(x_80); +lean_dec(x_79); +lean_dec(x_78); +lean_dec(x_77); +lean_dec(x_76); +lean_dec(x_75); lean_dec(x_74); lean_dec(x_73); -lean_dec(x_72); lean_dec(x_71); lean_dec(x_70); lean_dec(x_69); lean_dec(x_68); +lean_dec(x_67); lean_dec(x_66); -lean_dec(x_65); -lean_dec(x_64); -lean_dec(x_63); -lean_dec(x_62); -lean_dec(x_61); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4461,333 +6061,122 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_77 = lean_box(0); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_9); -return x_78; +x_87 = lean_box(0); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_9); +return x_88; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_79 = lean_ctor_get(x_76, 0); -lean_inc(x_79); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - x_80 = x_76; -} else { - lean_dec_ref(x_76); - x_80 = lean_box(0); -} -x_81 = lean_ctor_get(x_79, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_79, 1); -lean_inc(x_82); -lean_dec(x_79); -x_83 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_83, 0, x_61); -lean_ctor_set(x_83, 1, x_62); -lean_ctor_set(x_83, 2, x_63); -lean_ctor_set(x_83, 3, x_64); -lean_ctor_set(x_83, 4, x_65); -lean_ctor_set(x_83, 5, x_66); -lean_ctor_set(x_83, 6, x_68); -lean_ctor_set(x_83, 7, x_69); -lean_ctor_set(x_83, 8, x_70); -lean_ctor_set(x_83, 9, x_71); -lean_ctor_set(x_83, 10, x_72); -lean_ctor_set(x_83, 11, x_73); -lean_ctor_set(x_83, 12, x_74); -lean_ctor_set(x_83, 13, x_82); -lean_ctor_set_uint8(x_83, sizeof(void*)*14, x_67); -x_84 = lean_ctor_get(x_81, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_81, 1); -lean_inc(x_85); -x_86 = lean_ctor_get(x_81, 2); -lean_inc(x_86); -lean_dec(x_81); -x_87 = l_Lean_Meta_Grind_assertAt(x_83, x_84, x_85, x_86, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_87) == 0) -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_89 = lean_ctor_get(x_86, 0); 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; +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + x_90 = x_86; } else { - lean_dec_ref(x_87); + lean_dec_ref(x_86); x_90 = lean_box(0); } -if (lean_is_scalar(x_80)) { - x_91 = lean_alloc_ctor(1, 1, 0); -} else { - x_91 = x_80; -} -lean_ctor_set(x_91, 0, x_88); -if (lean_is_scalar(x_90)) { - x_92 = lean_alloc_ctor(0, 2, 0); -} else { - x_92 = x_90; -} -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_89); -return x_92; -} -else -{ -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_80); -x_93 = lean_ctor_get(x_87, 0); +x_91 = lean_ctor_get(x_89, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_89, 1); +lean_inc(x_92); +lean_dec(x_89); +x_93 = lean_ctor_get(x_91, 0); lean_inc(x_93); -x_94 = lean_ctor_get(x_87, 1); +x_94 = lean_ctor_get(x_91, 1); lean_inc(x_94); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - lean_ctor_release(x_87, 1); - x_95 = x_87; +x_95 = lean_ctor_get(x_91, 2); +lean_inc(x_95); +lean_dec(x_91); +x_96 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_96, 0, x_66); +lean_ctor_set(x_96, 1, x_67); +lean_ctor_set(x_96, 2, x_68); +lean_ctor_set(x_96, 3, x_69); +lean_ctor_set(x_96, 4, x_70); +lean_ctor_set(x_96, 5, x_71); +lean_ctor_set(x_96, 6, x_73); +lean_ctor_set(x_96, 7, x_74); +lean_ctor_set(x_96, 8, x_75); +lean_ctor_set(x_96, 9, x_76); +lean_ctor_set(x_96, 10, x_77); +lean_ctor_set(x_96, 11, x_78); +lean_ctor_set(x_96, 12, x_79); +lean_ctor_set(x_96, 13, x_80); +lean_ctor_set(x_96, 14, x_92); +lean_ctor_set(x_96, 15, x_82); +lean_ctor_set(x_96, 16, x_83); +lean_ctor_set(x_96, 17, x_84); +lean_ctor_set(x_96, 18, x_85); +lean_ctor_set_uint8(x_96, sizeof(void*)*19, x_72); +x_97 = l_Lean_Meta_Grind_assertAt(x_93, x_94, x_95, x_96, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_100 = x_97; } else { - lean_dec_ref(x_87); - x_95 = lean_box(0); + lean_dec_ref(x_97); + x_100 = lean_box(0); } -if (lean_is_scalar(x_95)) { - x_96 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_90)) { + x_101 = lean_alloc_ctor(1, 1, 0); } else { - x_96 = x_95; -} -lean_ctor_set(x_96, 0, x_93); -lean_ctor_set(x_96, 1, x_94); -return x_96; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_12; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_3); -lean_ctor_set(x_12, 1, x_11); -return x_12; -} -else -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_2); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_2, 0); -x_15 = lean_ctor_get(x_2, 1); -lean_inc(x_1); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_14); -x_16 = lean_apply_9(x_1, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -lean_ctor_set(x_2, 1, x_3); -{ -lean_object* _tmp_1 = x_15; -lean_object* _tmp_2 = x_2; -lean_object* _tmp_10 = x_18; -x_2 = _tmp_1; -x_3 = _tmp_2; -x_11 = _tmp_10; -} -goto _start; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_free_object(x_2); -lean_dec(x_14); -x_20 = lean_ctor_get(x_16, 1); -lean_inc(x_20); -lean_dec(x_16); -x_21 = lean_ctor_get(x_17, 0); -lean_inc(x_21); -lean_dec(x_17); -x_22 = l_List_appendTR___rarg(x_21, x_15); -x_2 = x_22; -x_11 = x_20; -goto _start; -} -} -else -{ -uint8_t x_24; -lean_free_object(x_2); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_16); -if (x_24 == 0) -{ -return x_16; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_16, 0); -x_26 = lean_ctor_get(x_16, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_16); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; -} -} -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_2, 0); -x_29 = lean_ctor_get(x_2, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_2); -lean_inc(x_1); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_28); -x_30 = lean_apply_9(x_1, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_28); -lean_ctor_set(x_33, 1, x_3); -x_2 = x_29; -x_3 = x_33; -x_11 = x_32; -goto _start; + x_101 = x_90; } -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -lean_dec(x_28); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_dec(x_30); -x_36 = lean_ctor_get(x_31, 0); -lean_inc(x_36); -lean_dec(x_31); -x_37 = l_List_appendTR___rarg(x_36, x_29); -x_2 = x_37; -x_11 = x_35; -goto _start; +lean_ctor_set(x_101, 0, x_98); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_100; } +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -lean_dec(x_29); -lean_dec(x_28); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_39 = lean_ctor_get(x_30, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_30, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_41 = x_30; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_90); +x_103 = lean_ctor_get(x_97, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_97, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_105 = x_97; } else { - lean_dec_ref(x_30); - x_41 = lean_box(0); + lean_dec_ref(x_97); + x_105 = lean_box(0); } -if (lean_is_scalar(x_41)) { - x_42 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_105)) { + x_106 = lean_alloc_ctor(1, 2, 0); } else { - x_42 = x_41; -} -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_40); -return x_42; -} + x_106 = x_105; } +lean_ctor_set(x_106, 0, x_103); +lean_ctor_set(x_106, 1, x_104); +return x_106; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_iterate(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Meta_Grind_iterate_go(x_2, x_12, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; } } static lean_object* _init_l_Lean_Meta_Grind_assertAll___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext_x3f), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_assertNext), 9, 0); return x_1; } } @@ -4796,7 +6185,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_assertAll(lean_object* x_1, lean_obje { lean_object* x_10; lean_object* x_11; x_10 = l_Lean_Meta_Grind_assertAll___closed__1; -x_11 = l_Lean_Meta_Grind_iterate(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Meta_Grind_GrindTactic_iterate(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } @@ -4807,6 +6196,7 @@ lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_objec lean_object* initialize_Lean_Meta_Tactic_Grind_Cases(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Injection(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Core(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Combinators(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object* w) { lean_object * res; @@ -4833,18 +6223,21 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Core(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Combinators(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Meta_Grind_instInhabitedIntroResult = _init_l_Lean_Meta_Grind_instInhabitedIntroResult(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedIntroResult); l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__1); -l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__2); -l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__3); -l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__5___closed__4); -l_Lean_Meta_Grind_intros_go___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_intros_go___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_intros_go___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__1); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__2); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__3); +l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Intro_0__Lean_Meta_Grind_introNext___lambda__8___closed__4); l_Lean_Meta_Grind_intros___closed__1 = _init_l_Lean_Meta_Grind_intros___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_intros___closed__1); l_Lean_Meta_Grind_assertAt___closed__1 = _init_l_Lean_Meta_Grind_assertAt___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c index dc2a5be0cb31..7b27280575ef 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Inv.c @@ -16,29 +16,30 @@ extern "C" { LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__6; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_checkInvariants___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___boxed(lean_object**); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__3; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Grind_checkInvariants___spec__3(lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3; uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); -lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5(lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -50,27 +51,23 @@ lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_grind_debug_proofs; -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1; -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6; lean_object* l_Lean_Meta_Grind_getEqcs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; +lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_grind_debug; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2; extern lean_object* l_instInhabitedNat; extern lean_object* l_instInhabitedPUnit; +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3; lean_object* l_Lean_Meta_Grind_getRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -81,41 +78,43 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_checkInvariants___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__2; static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getCongrRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__5; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__7; +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3; lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__4; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__9; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5; extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_checkInvariants___closed__1; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3___boxed(lean_object**); uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); uint8_t lean_expr_equal(lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); @@ -123,32 +122,33 @@ static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Me LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getParents(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__6; lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); lean_object* l_Lean_Meta_Grind_getENodes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1; +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___boxed(lean_object**); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__7; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5; +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Grind_checkInvariants___spec__5(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___closed__1; @@ -157,21 +157,20 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at___private_L lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5___closed__1; -lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__4; static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3; +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_isRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Grind_checkInvariants___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6; -lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkInvariants___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -181,39 +180,34 @@ size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__8___closed__1; lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntryAtAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2___closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__5; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3; -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_checkInvariants___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Grind_checkInvariants___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getRoot_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2; lean_object* l_ReaderT_instMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Grind_checkInvariants___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__2; -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__1; static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1___closed__1() { _start: @@ -636,6 +630,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findEntry_x3f___at___private_L lean_object* x_4; uint64_t x_5; size_t x_6; lean_object* x_7; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); +lean_inc(x_3); x_5 = l_Lean_Meta_Grind_congrHash(x_4, x_3); x_6 = lean_uint64_to_usize(x_5); x_7 = l_Lean_PersistentHashMap_findEntryAux___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__5(x_1, x_2, x_6, x_3); @@ -1092,7 +1087,7 @@ static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("isSameExpr curr ( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.34.0 ).cgRoot\n -- If the equivalence class does not have HEq proofs, then the types must be definitionally equal.\n ", 184, 184); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.34.0 )\n -- If the equivalence class does not have HEq proofs, then the types must be definitionally equal.\n ", 161, 161); return x_1; } } @@ -1123,7 +1118,7 @@ static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("isSameExpr e ( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.31.0 ).cgRoot\n ", 80, 80); +x_1 = lean_mk_string_unchecked("isSameExpr e ( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.31.0 )\n ", 73, 73); return x_1; } } @@ -1230,26 +1225,23 @@ if (lean_obj_tag(x_46) == 0) { lean_object* x_47; lean_inc(x_29); -x_47 = l_Lean_Meta_Grind_getENode(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); +x_47 = l_Lean_Meta_Grind_isCongrRoot(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); if (lean_obj_tag(x_47) == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +lean_object* x_48; uint8_t x_49; x_48 = lean_ctor_get(x_47, 0); lean_inc(x_48); -x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -lean_dec(x_47); -x_50 = lean_ctor_get(x_48, 3); -lean_inc(x_50); +x_49 = lean_unbox(x_48); lean_dec(x_48); -x_51 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_29, x_50); -lean_dec(x_50); -if (x_51 == 0) +if (x_49 == 0) { -lean_object* x_52; lean_object* x_53; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_dec(x_32); lean_dec(x_29); -x_52 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__6; +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); +lean_dec(x_47); +x_51 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__6; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -1258,13 +1250,16 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_53 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(x_52, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_49); -x_13 = x_53; +x_52 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(x_51, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_50); +x_13 = x_52; goto block_28; } else { -lean_object* x_54; lean_object* x_55; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_47, 1); +lean_inc(x_53); +lean_dec(x_47); x_54 = lean_box(0); lean_inc(x_11); lean_inc(x_10); @@ -1275,7 +1270,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(x_2, x_32, x_1, x_29, x_54, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_49); +x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(x_2, x_32, x_1, x_29, x_54, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_53); x_13 = x_55; goto block_28; } @@ -1358,27 +1353,24 @@ x_70 = lean_ctor_get(x_64, 1); lean_inc(x_70); lean_dec(x_64); lean_inc(x_29); -x_71 = l_Lean_Meta_Grind_getENode(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_70); +x_71 = l_Lean_Meta_Grind_getCongrRoot(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_70); if (lean_obj_tag(x_71) == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +lean_object* x_72; lean_object* x_73; uint8_t x_74; x_72 = lean_ctor_get(x_71, 0); lean_inc(x_72); x_73 = lean_ctor_get(x_71, 1); lean_inc(x_73); lean_dec(x_71); -x_74 = lean_ctor_get(x_72, 3); -lean_inc(x_74); +x_74 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_61, x_72); lean_dec(x_72); -x_75 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_61, x_74); -lean_dec(x_74); lean_dec(x_61); -if (x_75 == 0) +if (x_74 == 0) { -lean_object* x_76; lean_object* x_77; +lean_object* x_75; lean_object* x_76; lean_dec(x_32); lean_dec(x_29); -x_76 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__9; +x_75 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___closed__9; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -1387,14 +1379,14 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_77 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(x_76, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_73); -x_13 = x_77; +x_76 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__1(x_75, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_73); +x_13 = x_76; goto block_28; } else { -lean_object* x_78; lean_object* x_79; -x_78 = lean_box(0); +lean_object* x_77; lean_object* x_78; +x_77 = lean_box(0); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -1404,35 +1396,35 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -x_79 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(x_2, x_32, x_1, x_29, x_78, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_73); -x_13 = x_79; +x_78 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__2(x_2, x_32, x_1, x_29, x_77, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_73); +x_13 = x_78; goto block_28; } } else { -uint8_t x_80; +uint8_t x_79; lean_dec(x_61); lean_dec(x_32); lean_dec(x_29); -x_80 = !lean_is_exclusive(x_71); -if (x_80 == 0) +x_79 = !lean_is_exclusive(x_71); +if (x_79 == 0) { x_13 = x_71; goto block_28; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_81 = lean_ctor_get(x_71, 0); -x_82 = lean_ctor_get(x_71, 1); -lean_inc(x_82); +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_71, 0); +x_81 = lean_ctor_get(x_71, 1); lean_inc(x_81); +lean_inc(x_80); lean_dec(x_71); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -x_13 = x_83; +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +x_13 = x_82; goto block_28; } } @@ -1440,28 +1432,28 @@ goto block_28; } else { -uint8_t x_84; +uint8_t x_83; lean_dec(x_61); lean_dec(x_32); lean_dec(x_29); -x_84 = !lean_is_exclusive(x_64); -if (x_84 == 0) +x_83 = !lean_is_exclusive(x_64); +if (x_83 == 0) { x_13 = x_64; goto block_28; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_64, 0); -x_86 = lean_ctor_get(x_64, 1); -lean_inc(x_86); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_64, 0); +x_85 = lean_ctor_get(x_64, 1); lean_inc(x_85); +lean_inc(x_84); lean_dec(x_64); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -x_13 = x_87; +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_13 = x_86; goto block_28; } } @@ -1471,27 +1463,27 @@ goto block_28; } else { -uint8_t x_88; +uint8_t x_87; lean_dec(x_32); lean_dec(x_29); -x_88 = !lean_is_exclusive(x_33); -if (x_88 == 0) +x_87 = !lean_is_exclusive(x_33); +if (x_87 == 0) { x_13 = x_33; goto block_28; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_33, 0); -x_90 = lean_ctor_get(x_33, 1); -lean_inc(x_90); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_33, 0); +x_89 = lean_ctor_get(x_33, 1); lean_inc(x_89); +lean_inc(x_88); lean_dec(x_33); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_13 = x_91; +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +x_13 = x_90; goto block_28; } } @@ -1973,34 +1965,25 @@ x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_13; } } -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2() { +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.778.0 )\n ", 57, 57); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.782.0 )\n ", 57, 57); return x_1; } } -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3() { +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__1; -x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2; +x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4() { +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -2008,19 +1991,155 @@ x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Inv.0.Lean.Meta. return x_1; } } -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5() { +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; -x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4; -x_3 = lean_unsigned_to_nat(61u); +x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(65u); x_4 = lean_unsigned_to_nat(8u); -x_5 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3; +x_5 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +if (x_3 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = l_Lean_Expr_getAppFn(x_1); +x_15 = l_Lean_Meta_Grind_getRoot_x3f(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4; +x_19 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +return x_19; +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_ctor_get(x_15, 1); +x_22 = lean_ctor_get(x_15, 0); +lean_dec(x_22); +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +lean_dec(x_16); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_23, x_2); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_free_object(x_15); +x_25 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4; +x_26 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_21); +return x_26; +} +else +{ +lean_object* x_27; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +lean_ctor_set(x_15, 0, x_27); +return x_15; +} +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = lean_ctor_get(x_15, 1); +lean_inc(x_28); +lean_dec(x_15); +x_29 = lean_ctor_get(x_16, 0); +lean_inc(x_29); +lean_dec(x_16); +x_30 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_29, x_2); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4; +x_32 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_28); +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_33 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_28); +return x_34; +} +} +} +} +else +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_35 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_13); +return x_36; +} +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -2101,50 +2220,80 @@ return x_23; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +lean_object* x_24; lean_object* x_25; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; uint8_t x_49; lean_object* x_50; lean_dec(x_19); x_24 = lean_ctor_get(x_18, 1); lean_inc(x_24); lean_dec(x_18); -x_25 = lean_unsigned_to_nat(0u); -x_26 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_16, x_25); -x_27 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1; -lean_inc(x_26); -x_28 = lean_mk_array(x_26, x_27); -x_29 = lean_unsigned_to_nat(1u); -x_30 = lean_nat_sub(x_26, x_29); -lean_dec(x_26); +x_39 = lean_unsigned_to_nat(0u); +x_40 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_16, x_39); +x_41 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1; +lean_inc(x_40); +x_42 = lean_mk_array(x_40, x_41); +x_43 = lean_unsigned_to_nat(1u); +x_44 = lean_nat_sub(x_40, x_43); +lean_dec(x_40); lean_inc(x_16); -x_31 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_16, x_28, x_30); -x_32 = lean_box(0); -x_33 = lean_array_size(x_31); -x_34 = 0; -x_35 = 0; -x_36 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1(x_1, x_31, x_32, x_31, x_33, x_34, x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); -lean_dec(x_31); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_unbox(x_37); -lean_dec(x_37); -if (x_38 == 0) +x_45 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_16, x_42, x_44); +x_46 = lean_box(0); +x_47 = lean_array_size(x_45); +x_48 = 0; +x_49 = 0; +x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__1(x_1, x_45, x_46, x_45, x_47, x_48, x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +lean_dec(x_45); +if (lean_obj_tag(x_16) == 7) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_dec(x_36); -x_40 = l_Lean_Expr_getAppFn(x_16); +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_ctor_get(x_16, 1); +lean_inc(x_53); +x_54 = l_Lean_Meta_Grind_getRoot_x3f(x_53, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_52); +lean_dec(x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_box(0); +x_58 = lean_unbox(x_51); +lean_dec(x_51); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_59 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_16, x_1, x_58, x_57, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_56); lean_dec(x_16); -x_41 = l_Lean_Meta_Grind_getRoot_x3f(x_40, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); -lean_dec(x_40); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -if (lean_obj_tag(x_42) == 0) +x_25 = x_59; +goto block_38; +} +else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5; +lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_60 = lean_ctor_get(x_54, 1); +lean_inc(x_60); +lean_dec(x_54); +x_61 = lean_ctor_get(x_55, 0); +lean_inc(x_61); +lean_dec(x_55); +x_62 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_61, x_1); +lean_dec(x_61); +if (x_62 == 0) +{ +lean_object* x_63; uint8_t x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox(x_51); +lean_dec(x_51); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -2153,15 +2302,66 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_45 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_43); -if (lean_obj_tag(x_45) == 0) +x_65 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_16, x_1, x_64, x_63, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_60); +lean_dec(x_16); +x_25 = x_65; +goto block_38; +} +else { -lean_object* x_46; -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -if (lean_obj_tag(x_46) == 0) +uint8_t x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_51); +x_66 = 1; +x_67 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_68 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_16, x_1, x_66, x_67, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_60); +lean_dec(x_16); +x_25 = x_68; +goto block_38; +} +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; +x_69 = lean_ctor_get(x_50, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_50, 1); +lean_inc(x_70); +lean_dec(x_50); +x_71 = lean_box(0); +x_72 = lean_unbox(x_69); +lean_dec(x_69); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_73 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_16, x_1, x_72, x_71, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_70); +lean_dec(x_16); +x_25 = x_73; +goto block_38; +} +block_38: +{ +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +if (lean_obj_tag(x_26) == 0) { -uint8_t x_47; +uint8_t x_27; lean_dec(x_17); lean_dec(x_11); lean_dec(x_10); @@ -2171,44 +2371,44 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_47 = !lean_is_exclusive(x_45); -if (x_47 == 0) +x_27 = !lean_is_exclusive(x_25); +if (x_27 == 0) { -lean_object* x_48; -x_48 = lean_ctor_get(x_45, 0); -lean_dec(x_48); -return x_45; +lean_object* x_28; +x_28 = lean_ctor_get(x_25, 0); +lean_dec(x_28); +return x_25; } else { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_45, 1); -lean_inc(x_49); -lean_dec(x_45); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_46); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_25, 1); +lean_inc(x_29); +lean_dec(x_25); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_26); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } else { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_45, 1); -lean_inc(x_51); -lean_dec(x_45); -x_52 = lean_ctor_get(x_46, 0); -lean_inc(x_52); -lean_dec(x_46); +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); x_2 = x_17; -x_3 = x_52; -x_12 = x_51; +x_3 = x_32; +x_12 = x_31; goto _start; } } else { -uint8_t x_54; +uint8_t x_34; lean_dec(x_17); lean_dec(x_11); lean_dec(x_10); @@ -2218,163 +2418,31 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_54 = !lean_is_exclusive(x_45); -if (x_54 == 0) +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) { -return x_45; +return x_25; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_45, 0); -x_56 = lean_ctor_get(x_45, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_45); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; -} -} -} -else -{ -lean_object* x_58; lean_object* x_59; uint8_t x_60; -x_58 = lean_ctor_get(x_41, 1); -lean_inc(x_58); -lean_dec(x_41); -x_59 = lean_ctor_get(x_42, 0); -lean_inc(x_59); -lean_dec(x_42); -x_60 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_59, x_1); -lean_dec(x_59); -if (x_60 == 0) -{ -lean_object* x_61; lean_object* x_62; -x_61 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_62 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_61, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_58); -if (lean_obj_tag(x_62) == 0) -{ -lean_object* x_63; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -if (lean_obj_tag(x_63) == 0) -{ -uint8_t x_64; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_64 = !lean_is_exclusive(x_62); -if (x_64 == 0) -{ -lean_object* x_65; -x_65 = lean_ctor_get(x_62, 0); -lean_dec(x_65); -return x_62; -} -else -{ -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_62, 1); -lean_inc(x_66); -lean_dec(x_62); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_63); -lean_ctor_set(x_67, 1, x_66); -return x_67; -} -} -else -{ -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_62, 1); -lean_inc(x_68); -lean_dec(x_62); -x_69 = lean_ctor_get(x_63, 0); -lean_inc(x_69); -lean_dec(x_63); -x_2 = x_17; -x_3 = x_69; -x_12 = x_68; -goto _start; -} -} -else -{ -uint8_t x_71; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_71 = !lean_is_exclusive(x_62); -if (x_71 == 0) -{ -return x_62; -} -else -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_62, 0); -x_73 = lean_ctor_get(x_62, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_62); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; -} -} -} -else -{ -lean_object* x_75; -x_75 = lean_box(0); -x_2 = x_17; -x_3 = x_75; -x_12 = x_58; -goto _start; -} +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_ctor_get(x_25, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_25); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } -else -{ -lean_object* x_77; lean_object* x_78; -lean_dec(x_16); -x_77 = lean_ctor_get(x_36, 1); -lean_inc(x_77); -lean_dec(x_36); -x_78 = lean_box(0); -x_2 = x_17; -x_3 = x_78; -x_12 = x_77; -goto _start; } } } else { -uint8_t x_80; +uint8_t x_74; lean_dec(x_17); lean_dec(x_16); lean_dec(x_11); @@ -2385,23 +2453,23 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_80 = !lean_is_exclusive(x_18); -if (x_80 == 0) +x_74 = !lean_is_exclusive(x_18); +if (x_74 == 0) { return x_18; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_81 = lean_ctor_get(x_18, 0); -x_82 = lean_ctor_get(x_18, 1); -lean_inc(x_82); -lean_inc(x_81); +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_18, 0); +x_76 = lean_ctor_get(x_18, 1); +lean_inc(x_76); +lean_inc(x_75); lean_dec(x_18); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -return x_83; +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +return x_77; } } } @@ -2411,7 +2479,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Gr _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.791.0 ).isEmpty\n\n", 64, 64); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Inv._hyg.805.0 ).isEmpty\n\n", 64, 64); return x_1; } } @@ -2430,8 +2498,8 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Gr { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; -x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4; -x_3 = lean_unsigned_to_nat(64u); +x_2 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(68u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2591,6 +2659,19 @@ lean_dec(x_1); return x_20; } } +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_3); +lean_dec(x_3); +x_15 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -2609,7 +2690,17 @@ lean_dec(x_1); return x_11; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1() { _start: { lean_object* x_1; @@ -2617,17 +2708,17 @@ x_1 = lean_mk_string_unchecked("!Expr.equal n₁.self n₂.self\n\n", 33, 29); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__1; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1; +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3() { _start: { lean_object* x_1; @@ -2635,20 +2726,20 @@ x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Inv.0.Lean.Meta. return x_1; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3; -x_3 = lean_unsigned_to_nat(75u); +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3; +x_3 = lean_unsigned_to_nat(79u); x_4 = lean_unsigned_to_nat(6u); -x_5 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2; +x_5 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5() { _start: { lean_object* x_1; @@ -2656,30 +2747,30 @@ x_1 = lean_mk_string_unchecked("!isSameExpr n₁.self n₂.self\n -- and th return x_1; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__1; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5; +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7() { +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkEqc___spec__7___lambda__1___closed__4; -x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3; -x_3 = lean_unsigned_to_nat(73u); +x_2 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3; +x_3 = lean_unsigned_to_nat(77u); x_4 = lean_unsigned_to_nat(6u); -x_5 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6; +x_5 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_18; uint8_t x_19; @@ -2733,7 +2824,7 @@ goto _start; else { lean_object* x_30; lean_object* x_31; -x_30 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4; +x_30 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4; lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); @@ -2742,7 +2833,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_31 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_30, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_31 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(x_30, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; @@ -2844,7 +2935,7 @@ else { lean_object* x_48; lean_object* x_49; lean_dec(x_23); -x_48 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7; +x_48 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7; lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); @@ -2853,7 +2944,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_49 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2(x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_49 = l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(x_48, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); if (lean_obj_tag(x_49) == 0) { lean_object* x_50; @@ -2953,7 +3044,7 @@ return x_65; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_18; uint8_t x_19; @@ -2999,7 +3090,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_26 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(x_1, x_21, x_24, x_24, x_25, x_23, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_26 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(x_1, x_21, x_24, x_24, x_25, x_23, lean_box(0), lean_box(0), x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_24); lean_dec(x_21); if (lean_obj_tag(x_26) == 0) @@ -3072,7 +3163,7 @@ lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_13); lean_ctor_set(x_16, 2, x_15); x_17 = lean_box(0); -x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(x_11, x_13, x_16, x_16, x_17, x_14, lean_box(0), lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3(x_11, x_13, x_16, x_16, x_17, x_14, lean_box(0), lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); lean_dec(x_16); lean_dec(x_11); if (lean_obj_tag(x_18) == 0) @@ -3123,7 +3214,7 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -3144,7 +3235,7 @@ lean_object* x_17 = _args[16]; _start: { lean_object* x_18; -x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -3152,7 +3243,7 @@ lean_dec(x_1); return x_18; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -3173,50 +3264,31 @@ lean_object* x_17 = _args[16]; _start: { lean_object* x_18; -x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); return x_18; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1() { _start: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 2); -x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -return x_12; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("checked: ", 9, 9); +return x_1; } } -static double _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2() { _start: { -lean_object* x_1; uint8_t x_2; double x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = 0; -x_3 = l_Float_ofScientific(x_1, x_2, x_1); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -3224,539 +3296,221 @@ x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } -static lean_object* _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_12 = lean_ctor_get(x_9, 5); -x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); -x_14 = lean_ctor_get(x_13, 0); +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_13 = l_Lean_Meta_check(x_1, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); lean_dec(x_13); -x_16 = lean_st_ref_take(x_10, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 3); -lean_inc(x_18); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_16, 1); -x_21 = lean_ctor_get(x_16, 0); -lean_dec(x_21); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_17, 3); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_18); -if (x_24 == 0) +lean_inc(x_2); +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) { -lean_object* x_25; double x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_25 = lean_ctor_get(x_18, 0); -x_26 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; -x_27 = 0; -x_28 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_29 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); -lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); -lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); -x_30 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; -x_31 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_14); -lean_ctor_set(x_31, 2, x_30); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_31); -lean_ctor_set(x_16, 0, x_12); -x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); -lean_ctor_set(x_18, 0, x_32); -x_33 = lean_st_ref_set(x_10, x_17, x_20); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +uint8_t x_18; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_33, 0, x_36); -return x_33; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_15, 0); +lean_dec(x_19); +x_20 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +lean_ctor_set(x_15, 0, x_20); +return x_15; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +x_22 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; } } else { -uint64_t x_40; lean_object* x_41; double x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_41 = lean_ctor_get(x_18, 0); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; -x_43 = 0; -x_44 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_45 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); -lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); -x_46 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; -x_47 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_46); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_47); -lean_ctor_set(x_16, 0, x_12); -x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); -x_49 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); -lean_ctor_set(x_17, 3, x_49); -x_50 = lean_st_ref_set(x_10, x_17, x_20); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); -} -x_53 = lean_box(0); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_51); -return x_54; -} +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_15, 1); +lean_inc(x_24); +lean_dec(x_15); +x_25 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_27 = lean_infer_type(x_1, x_8, x_9, x_10, x_11, x_26); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_MessageData_ofExpr(x_28); +x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_2, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_29); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +lean_ctor_set(x_35, 0, x_38); +return x_35; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_55 = lean_ctor_get(x_17, 0); -x_56 = lean_ctor_get(x_17, 1); -x_57 = lean_ctor_get(x_17, 2); -x_58 = lean_ctor_get(x_17, 4); -x_59 = lean_ctor_get(x_17, 5); -x_60 = lean_ctor_get(x_17, 6); -x_61 = lean_ctor_get(x_17, 7); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_17); -x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_63 = lean_ctor_get(x_18, 0); -lean_inc(x_63); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_64 = x_18; -} else { - lean_dec_ref(x_18); - x_64 = lean_box(0); -} -x_65 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; -x_66 = 0; -x_67 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_68 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_68, 0, x_1); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); -lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); -lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); -x_69 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; -x_70 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_14); -lean_ctor_set(x_70, 2, x_69); -lean_inc(x_12); -lean_ctor_set(x_16, 1, x_70); -lean_ctor_set(x_16, 0, x_12); -x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); -if (lean_is_scalar(x_64)) { - x_72 = lean_alloc_ctor(0, 1, 8); -} else { - x_72 = x_64; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); -x_73 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_73, 0, x_55); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_57); -lean_ctor_set(x_73, 3, x_72); -lean_ctor_set(x_73, 4, x_58); -lean_ctor_set(x_73, 5, x_59); -lean_ctor_set(x_73, 6, x_60); -lean_ctor_set(x_73, 7, x_61); -x_74 = lean_st_ref_set(x_10, x_73, x_20); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); -} -x_77 = lean_box(0); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_75); -return x_78; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_79 = lean_ctor_get(x_16, 1); -lean_inc(x_79); -lean_dec(x_16); -x_80 = lean_ctor_get(x_17, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_17, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_17, 2); -lean_inc(x_82); -x_83 = lean_ctor_get(x_17, 4); -lean_inc(x_83); -x_84 = lean_ctor_get(x_17, 5); -lean_inc(x_84); -x_85 = lean_ctor_get(x_17, 6); -lean_inc(x_85); -x_86 = lean_ctor_get(x_17, 7); -lean_inc(x_86); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - lean_ctor_release(x_17, 2); - lean_ctor_release(x_17, 3); - lean_ctor_release(x_17, 4); - lean_ctor_release(x_17, 5); - lean_ctor_release(x_17, 6); - lean_ctor_release(x_17, 7); - x_87 = x_17; -} else { - lean_dec_ref(x_17); - x_87 = lean_box(0); -} -x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); -x_89 = lean_ctor_get(x_18, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_18)) { - lean_ctor_release(x_18, 0); - x_90 = x_18; -} else { - lean_dec_ref(x_18); - x_90 = lean_box(0); -} -x_91 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1; -x_92 = 0; -x_93 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_94 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); -lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); -x_95 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3; -x_96 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_14); -lean_ctor_set(x_96, 2, x_95); -lean_inc(x_12); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_12); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); -if (lean_is_scalar(x_90)) { - x_99 = lean_alloc_ctor(0, 1, 8); -} else { - x_99 = x_90; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); -if (lean_is_scalar(x_87)) { - x_100 = lean_alloc_ctor(0, 8, 0); -} else { - x_100 = x_87; -} -lean_ctor_set(x_100, 0, x_80); -lean_ctor_set(x_100, 1, x_81); -lean_ctor_set(x_100, 2, x_82); -lean_ctor_set(x_100, 3, x_99); -lean_ctor_set(x_100, 4, x_83); -lean_ctor_set(x_100, 5, x_84); -lean_ctor_set(x_100, 6, x_85); -lean_ctor_set(x_100, 7, x_86); -x_101 = lean_st_ref_set(x_10, x_100, x_79); -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_103 = x_101; -} else { - lean_dec_ref(x_101); - x_103 = lean_box(0); -} -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; -} -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("checked: ", 9, 9); -return x_1; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_13 = l_Lean_Meta_check(x_1, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -lean_inc(x_2); -x_15 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) -{ -uint8_t x_18; +uint8_t x_42; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_2); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_15, 0); -lean_dec(x_19); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; -lean_ctor_set(x_15, 0, x_20); -return x_15; -} -else +x_42 = !lean_is_exclusive(x_27); +if (x_42 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); -lean_dec(x_15); -x_22 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} +return x_27; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_15, 1); -lean_inc(x_24); -lean_dec(x_15); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_25 = lean_infer_type(x_1, x_8, x_9, x_10, x_11, x_24); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = l_Lean_MessageData_ofExpr(x_26); -x_29 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_2, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_dec(x_35); -x_36 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; -lean_ctor_set(x_33, 0, x_36); -return x_33; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_27, 0); +x_44 = lean_ctor_get(x_27, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_27); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; } } else { -uint8_t x_40; +uint8_t x_46; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_2); -x_40 = !lean_is_exclusive(x_25); -if (x_40 == 0) +lean_dec(x_1); +x_46 = !lean_is_exclusive(x_25); +if (x_46 == 0) { return x_25; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_25, 0); -x_42 = lean_ctor_get(x_25, 1); -lean_inc(x_42); -lean_inc(x_41); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_25, 0); +x_48 = lean_ctor_get(x_25, 1); +lean_inc(x_48); +lean_inc(x_47); lean_dec(x_25); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } } } else { -uint8_t x_44; +uint8_t x_50; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_2); lean_dec(x_1); -x_44 = !lean_is_exclusive(x_13); -if (x_44 == 0) +x_50 = !lean_is_exclusive(x_13); +if (x_50 == 0) { return x_13; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_13, 0); -x_46 = lean_ctor_get(x_13, 1); -lean_inc(x_46); -lean_inc(x_45); +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_13, 0); +x_52 = lean_ctor_get(x_13, 1); +lean_inc(x_52); +lean_inc(x_51); lean_dec(x_13); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1() { _start: { lean_object* x_1; @@ -3764,7 +3518,7 @@ x_1 = lean_mk_string_unchecked("grind", 5, 5); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2() { _start: { lean_object* x_1; @@ -3772,7 +3526,7 @@ x_1 = lean_mk_string_unchecked("debug", 5, 5); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3() { _start: { lean_object* x_1; @@ -3780,18 +3534,18 @@ x_1 = lean_mk_string_unchecked("proofs", 6, 6); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1; -x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2; -x_3 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1; +x_2 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2; +x_3 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5() { _start: { lean_object* x_1; @@ -3799,16 +3553,16 @@ x_1 = lean_mk_string_unchecked(" = ", 3, 3); return x_1; } } -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6() { +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5; +x_1 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { if (lean_obj_tag(x_5) == 0) @@ -3861,8 +3615,8 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; -x_26 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +x_25 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; +x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_24); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_unbox(x_27); @@ -3880,7 +3634,7 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +x_31 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_29); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; @@ -3981,47 +3735,54 @@ uint8_t x_46; x_46 = !lean_is_exclusive(x_26); if (x_46 == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_47; lean_object* x_48; lean_object* x_49; x_47 = lean_ctor_get(x_26, 1); x_48 = lean_ctor_get(x_26, 0); lean_dec(x_48); +x_49 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_47); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); lean_inc(x_3); -x_49 = l_Lean_MessageData_ofExpr(x_3); -x_50 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; +x_51 = l_Lean_MessageData_ofExpr(x_3); +x_52 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; lean_ctor_set_tag(x_26, 7); -lean_ctor_set(x_26, 1, x_49); -lean_ctor_set(x_26, 0, x_50); -x_51 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; +lean_ctor_set(x_26, 1, x_51); +lean_ctor_set(x_26, 0, x_52); +x_53 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; lean_ctor_set_tag(x_5, 7); -lean_ctor_set(x_5, 1, x_51); +lean_ctor_set(x_5, 1, x_53); lean_ctor_set(x_5, 0, x_26); -x_52 = l_Lean_MessageData_ofExpr(x_19); -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_5); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_50); -x_55 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_25, x_54, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_47); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); +x_54 = l_Lean_MessageData_ofExpr(x_19); +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_5); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_52); +x_57 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_50); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_58 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_57); -lean_dec(x_56); -if (lean_obj_tag(x_58) == 0) +x_60 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_58, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_59); +lean_dec(x_58); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_59; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -if (lean_obj_tag(x_59) == 0) +lean_object* x_61; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +if (lean_obj_tag(x_61) == 0) { -uint8_t x_60; +uint8_t x_62; lean_dec(x_20); lean_dec(x_15); lean_dec(x_14); @@ -4032,52 +3793,52 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_60 = !lean_is_exclusive(x_58); -if (x_60 == 0) +x_62 = !lean_is_exclusive(x_60); +if (x_62 == 0) { -lean_object* x_61; lean_object* x_62; -x_61 = lean_ctor_get(x_58, 0); +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_60, 0); +lean_dec(x_63); +x_64 = lean_ctor_get(x_61, 0); +lean_inc(x_64); lean_dec(x_61); -x_62 = lean_ctor_get(x_59, 0); -lean_inc(x_62); -lean_dec(x_59); -lean_ctor_set(x_58, 0, x_62); -return x_58; +lean_ctor_set(x_60, 0, x_64); +return x_60; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_58, 1); -lean_inc(x_63); -lean_dec(x_58); -x_64 = lean_ctor_get(x_59, 0); -lean_inc(x_64); -lean_dec(x_59); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_63); -return x_65; +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_60, 1); +lean_inc(x_65); +lean_dec(x_60); +x_66 = lean_ctor_get(x_61, 0); +lean_inc(x_66); +lean_dec(x_61); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_65); +return x_67; } } else { -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_58, 1); -lean_inc(x_66); -lean_dec(x_58); -x_67 = lean_ctor_get(x_59, 0); -lean_inc(x_67); -lean_dec(x_59); +lean_object* x_68; lean_object* x_69; +x_68 = lean_ctor_get(x_60, 1); +lean_inc(x_68); +lean_dec(x_60); +x_69 = lean_ctor_get(x_61, 0); +lean_inc(x_69); +lean_dec(x_61); x_5 = x_20; -x_6 = x_67; +x_6 = x_69; x_7 = lean_box(0); -x_16 = x_66; +x_16 = x_68; goto _start; } } else { -uint8_t x_69; +uint8_t x_71; lean_dec(x_20); lean_dec(x_15); lean_dec(x_14); @@ -4088,69 +3849,113 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_69 = !lean_is_exclusive(x_58); -if (x_69 == 0) +x_71 = !lean_is_exclusive(x_60); +if (x_71 == 0) { -return x_58; +return x_60; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_58, 0); -x_71 = lean_ctor_get(x_58, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_58); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_60, 0); +x_73 = lean_ctor_get(x_60, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_60); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; } } } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_73 = lean_ctor_get(x_26, 1); -lean_inc(x_73); +uint8_t x_75; +lean_free_object(x_26); +lean_dec(x_23); +lean_free_object(x_5); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_75 = !lean_is_exclusive(x_49); +if (x_75 == 0) +{ +return x_49; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_49, 0); +x_77 = lean_ctor_get(x_49, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_49); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; +x_79 = lean_ctor_get(x_26, 1); +lean_inc(x_79); lean_dec(x_26); +x_80 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_79); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); lean_inc(x_3); -x_74 = l_Lean_MessageData_ofExpr(x_3); -x_75 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; -x_76 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_74); -x_77 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; +x_82 = l_Lean_MessageData_ofExpr(x_3); +x_83 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_82); +x_85 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; lean_ctor_set_tag(x_5, 7); -lean_ctor_set(x_5, 1, x_77); -lean_ctor_set(x_5, 0, x_76); -x_78 = l_Lean_MessageData_ofExpr(x_19); -x_79 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_79, 0, x_5); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_75); -x_81 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_25, x_80, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_73); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); +lean_ctor_set(x_5, 1, x_85); +lean_ctor_set(x_5, 0, x_84); +x_86 = l_Lean_MessageData_ofExpr(x_19); +x_87 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_87, 0, x_5); +lean_ctor_set(x_87, 1, x_86); +x_88 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_83); +x_89 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_25, x_88, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_81); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_84 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_23, x_25, x_82, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_83); -lean_dec(x_82); -if (lean_obj_tag(x_84) == 0) +x_92 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_23, x_25, x_90, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_91); +lean_dec(x_90); +if (lean_obj_tag(x_92) == 0) { -lean_object* x_85; -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -if (lean_obj_tag(x_85) == 0) +lean_object* x_93; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +if (lean_obj_tag(x_93) == 0) { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_dec(x_20); lean_dec(x_15); lean_dec(x_14); @@ -4161,47 +3966,47 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_86 = lean_ctor_get(x_84, 1); -lean_inc(x_86); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_87 = x_84; +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_95 = x_92; } else { - lean_dec_ref(x_84); - x_87 = lean_box(0); + lean_dec_ref(x_92); + x_95 = lean_box(0); } -x_88 = lean_ctor_get(x_85, 0); -lean_inc(x_88); -lean_dec(x_85); -if (lean_is_scalar(x_87)) { - x_89 = lean_alloc_ctor(0, 2, 0); +x_96 = lean_ctor_get(x_93, 0); +lean_inc(x_96); +lean_dec(x_93); +if (lean_is_scalar(x_95)) { + x_97 = lean_alloc_ctor(0, 2, 0); } else { - x_89 = x_87; + x_97 = x_95; } -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_86); -return x_89; +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_94); +return x_97; } else { -lean_object* x_90; lean_object* x_91; -x_90 = lean_ctor_get(x_84, 1); -lean_inc(x_90); -lean_dec(x_84); -x_91 = lean_ctor_get(x_85, 0); -lean_inc(x_91); -lean_dec(x_85); +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_92, 1); +lean_inc(x_98); +lean_dec(x_92); +x_99 = lean_ctor_get(x_93, 0); +lean_inc(x_99); +lean_dec(x_93); x_5 = x_20; -x_6 = x_91; +x_6 = x_99; x_7 = lean_box(0); -x_16 = x_90; +x_16 = x_98; goto _start; } } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_dec(x_20); lean_dec(x_15); lean_dec(x_14); @@ -4212,33 +4017,71 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_93 = lean_ctor_get(x_84, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_84, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_95 = x_84; +x_101 = lean_ctor_get(x_92, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_92, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_103 = x_92; } else { - lean_dec_ref(x_84); - x_95 = lean_box(0); + lean_dec_ref(x_92); + x_103 = lean_box(0); } -if (lean_is_scalar(x_95)) { - x_96 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(1, 2, 0); } else { - x_96 = x_95; + x_104 = x_103; } -lean_ctor_set(x_96, 0, x_93); -lean_ctor_set(x_96, 1, x_94); -return x_96; +lean_ctor_set(x_104, 0, x_101); +lean_ctor_set(x_104, 1, x_102); +return x_104; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_dec(x_23); +lean_free_object(x_5); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_105 = lean_ctor_get(x_80, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_80, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_107 = x_80; +} else { + lean_dec_ref(x_80); + x_107 = lean_box(0); +} +if (lean_is_scalar(x_107)) { + x_108 = lean_alloc_ctor(1, 2, 0); +} else { + x_108 = x_107; +} +lean_ctor_set(x_108, 0, x_105); +lean_ctor_set(x_108, 1, x_106); +return x_108; } } } } else { -uint8_t x_97; +uint8_t x_109; lean_free_object(x_5); lean_dec(x_20); lean_dec(x_19); @@ -4251,50 +4094,50 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_97 = !lean_is_exclusive(x_22); -if (x_97 == 0) +x_109 = !lean_is_exclusive(x_22); +if (x_109 == 0) { return x_22; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_22, 0); -x_99 = lean_ctor_get(x_22, 1); -lean_inc(x_99); -lean_inc(x_98); +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_22, 0); +x_111 = lean_ctor_get(x_22, 1); +lean_inc(x_111); +lean_inc(x_110); lean_dec(x_22); -x_100 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -return x_100; +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } else { -lean_object* x_101; +lean_object* x_113; lean_free_object(x_5); lean_dec(x_19); -x_101 = lean_box(0); +x_113 = lean_box(0); x_5 = x_20; -x_6 = x_101; +x_6 = x_113; x_7 = lean_box(0); goto _start; } } else { -lean_object* x_103; lean_object* x_104; uint8_t x_105; -x_103 = lean_ctor_get(x_5, 0); -x_104 = lean_ctor_get(x_5, 1); -lean_inc(x_104); -lean_inc(x_103); +lean_object* x_115; lean_object* x_116; uint8_t x_117; +x_115 = lean_ctor_get(x_5, 0); +x_116 = lean_ctor_get(x_5, 1); +lean_inc(x_116); +lean_inc(x_115); lean_dec(x_5); -x_105 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_103); -if (x_105 == 0) +x_117 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_115); +if (x_117 == 0) { -lean_object* x_106; +lean_object* x_118; lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); @@ -4303,45 +4146,45 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_103); +lean_inc(x_115); lean_inc(x_3); -x_106 = l_Lean_Meta_Grind_mkEqHEqProof(x_3, x_103, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_106) == 0) -{ -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); -lean_dec(x_106); -x_109 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4; -x_110 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_109, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_108); -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -x_112 = lean_unbox(x_111); -lean_dec(x_111); -if (x_112 == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; -lean_dec(x_103); -x_113 = lean_ctor_get(x_110, 1); -lean_inc(x_113); -lean_dec(x_110); -x_114 = lean_box(0); +x_118 = l_Lean_Meta_Grind_mkEqHEqProof(x_3, x_115, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_118) == 0) +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4; +x_122 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_121, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_120); +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_unbox(x_123); +lean_dec(x_123); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_115); +x_125 = lean_ctor_get(x_122, 1); +lean_inc(x_125); +lean_dec(x_122); +x_126 = lean_box(0); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_115 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_107, x_109, x_114, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_113); -if (lean_obj_tag(x_115) == 0) +x_127 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_119, x_121, x_126, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_125); +if (lean_obj_tag(x_127) == 0) { -lean_object* x_116; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -if (lean_obj_tag(x_116) == 0) +lean_object* x_128; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +if (lean_obj_tag(x_128) == 0) { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -lean_dec(x_104); +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +lean_dec(x_116); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4351,48 +4194,48 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_117 = lean_ctor_get(x_115, 1); -lean_inc(x_117); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - x_118 = x_115; +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_130 = x_127; } else { - lean_dec_ref(x_115); - x_118 = lean_box(0); -} -x_119 = lean_ctor_get(x_116, 0); -lean_inc(x_119); -lean_dec(x_116); -if (lean_is_scalar(x_118)) { - x_120 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_127); + x_130 = lean_box(0); +} +x_131 = lean_ctor_get(x_128, 0); +lean_inc(x_131); +lean_dec(x_128); +if (lean_is_scalar(x_130)) { + x_132 = lean_alloc_ctor(0, 2, 0); } else { - x_120 = x_118; + x_132 = x_130; } -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_117); -return x_120; +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_129); +return x_132; } else { -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_115, 1); -lean_inc(x_121); -lean_dec(x_115); -x_122 = lean_ctor_get(x_116, 0); -lean_inc(x_122); -lean_dec(x_116); -x_5 = x_104; -x_6 = x_122; +lean_object* x_133; lean_object* x_134; +x_133 = lean_ctor_get(x_127, 1); +lean_inc(x_133); +lean_dec(x_127); +x_134 = lean_ctor_get(x_128, 0); +lean_inc(x_134); +lean_dec(x_128); +x_5 = x_116; +x_6 = x_134; x_7 = lean_box(0); -x_16 = x_121; +x_16 = x_133; goto _start; } } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_104); +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_116); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4402,84 +4245,91 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_124 = lean_ctor_get(x_115, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_115, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - x_126 = x_115; +x_136 = lean_ctor_get(x_127, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_127, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_138 = x_127; } else { - lean_dec_ref(x_115); - x_126 = lean_box(0); + lean_dec_ref(x_127); + x_138 = lean_box(0); } -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_138)) { + x_139 = lean_alloc_ctor(1, 2, 0); } else { - x_127 = x_126; + x_139 = x_138; } -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_ctor_set(x_139, 0, x_136); +lean_ctor_set(x_139, 1, x_137); +return x_139; } } else { -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_128 = lean_ctor_get(x_110, 1); -lean_inc(x_128); -if (lean_is_exclusive(x_110)) { - lean_ctor_release(x_110, 0); - lean_ctor_release(x_110, 1); - x_129 = x_110; +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_122, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_141 = x_122; } else { - lean_dec_ref(x_110); - x_129 = lean_box(0); + lean_dec_ref(x_122); + x_141 = lean_box(0); } +x_142 = l_Lean_Meta_Grind_updateLastTag(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_140); +if (lean_obj_tag(x_142) == 0) +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_143 = lean_ctor_get(x_142, 1); +lean_inc(x_143); +lean_dec(x_142); lean_inc(x_3); -x_130 = l_Lean_MessageData_ofExpr(x_3); -x_131 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4; -if (lean_is_scalar(x_129)) { - x_132 = lean_alloc_ctor(7, 2, 0); +x_144 = l_Lean_MessageData_ofExpr(x_3); +x_145 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4; +if (lean_is_scalar(x_141)) { + x_146 = lean_alloc_ctor(7, 2, 0); } else { - x_132 = x_129; - lean_ctor_set_tag(x_132, 7); + x_146 = x_141; + lean_ctor_set_tag(x_146, 7); } -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_130); -x_133 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6; -x_134 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_134, 0, x_132); -lean_ctor_set(x_134, 1, x_133); -x_135 = l_Lean_MessageData_ofExpr(x_103); -x_136 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_135); -x_137 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_131); -x_138 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_109, x_137, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_128); -x_139 = lean_ctor_get(x_138, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_138, 1); -lean_inc(x_140); -lean_dec(x_138); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +x_147 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6; +x_148 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +x_149 = l_Lean_MessageData_ofExpr(x_115); +x_150 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +x_151 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_145); +x_152 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_121, x_151, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_143); +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +lean_dec(x_152); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_141 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_107, x_109, x_139, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_140); -lean_dec(x_139); -if (lean_obj_tag(x_141) == 0) +x_155 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_119, x_121, x_153, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_154); +lean_dec(x_153); +if (lean_obj_tag(x_155) == 0) { -lean_object* x_142; -x_142 = lean_ctor_get(x_141, 0); -lean_inc(x_142); -if (lean_obj_tag(x_142) == 0) +lean_object* x_156; +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +if (lean_obj_tag(x_156) == 0) { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_104); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +lean_dec(x_116); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4489,48 +4339,86 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_143 = lean_ctor_get(x_141, 1); -lean_inc(x_143); -if (lean_is_exclusive(x_141)) { - lean_ctor_release(x_141, 0); - lean_ctor_release(x_141, 1); - x_144 = x_141; +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_158 = x_155; } else { - lean_dec_ref(x_141); - x_144 = lean_box(0); -} -x_145 = lean_ctor_get(x_142, 0); -lean_inc(x_145); -lean_dec(x_142); -if (lean_is_scalar(x_144)) { - x_146 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_155); + x_158 = lean_box(0); +} +x_159 = lean_ctor_get(x_156, 0); +lean_inc(x_159); +lean_dec(x_156); +if (lean_is_scalar(x_158)) { + x_160 = lean_alloc_ctor(0, 2, 0); } else { - x_146 = x_144; + x_160 = x_158; } -lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_146, 1, x_143); -return x_146; +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_157); +return x_160; } else { -lean_object* x_147; lean_object* x_148; -x_147 = lean_ctor_get(x_141, 1); -lean_inc(x_147); -lean_dec(x_141); -x_148 = lean_ctor_get(x_142, 0); -lean_inc(x_148); -lean_dec(x_142); -x_5 = x_104; -x_6 = x_148; +lean_object* x_161; lean_object* x_162; +x_161 = lean_ctor_get(x_155, 1); +lean_inc(x_161); +lean_dec(x_155); +x_162 = lean_ctor_get(x_156, 0); +lean_inc(x_162); +lean_dec(x_156); +x_5 = x_116; +x_6 = x_162; x_7 = lean_box(0); -x_16 = x_147; +x_16 = x_161; goto _start; } } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -lean_dec(x_104); +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +lean_dec(x_116); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_164 = lean_ctor_get(x_155, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_155, 1); +lean_inc(x_165); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_166 = x_155; +} else { + lean_dec_ref(x_155); + x_166 = lean_box(0); +} +if (lean_is_scalar(x_166)) { + x_167 = lean_alloc_ctor(1, 2, 0); +} else { + x_167 = x_166; +} +lean_ctor_set(x_167, 0, x_164); +lean_ctor_set(x_167, 1, x_165); +return x_167; +} +} +else +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +lean_dec(x_141); +lean_dec(x_119); +lean_dec(x_116); +lean_dec(x_115); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4540,34 +4428,34 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_150 = lean_ctor_get(x_141, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_141, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_141)) { - lean_ctor_release(x_141, 0); - lean_ctor_release(x_141, 1); - x_152 = x_141; +x_168 = lean_ctor_get(x_142, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_142, 1); +lean_inc(x_169); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + x_170 = x_142; } else { - lean_dec_ref(x_141); - x_152 = lean_box(0); + lean_dec_ref(x_142); + x_170 = lean_box(0); } -if (lean_is_scalar(x_152)) { - x_153 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_170)) { + x_171 = lean_alloc_ctor(1, 2, 0); } else { - x_153 = x_152; + x_171 = x_170; } -lean_ctor_set(x_153, 0, x_150); -lean_ctor_set(x_153, 1, x_151); -return x_153; +lean_ctor_set(x_171, 0, x_168); +lean_ctor_set(x_171, 1, x_169); +return x_171; } } } else { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -lean_dec(x_104); -lean_dec(x_103); +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_116); +lean_dec(x_115); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -4577,35 +4465,35 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_3); -x_154 = lean_ctor_get(x_106, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_106, 1); -lean_inc(x_155); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_156 = x_106; +x_172 = lean_ctor_get(x_118, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_118, 1); +lean_inc(x_173); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_174 = x_118; } else { - lean_dec_ref(x_106); - x_156 = lean_box(0); + lean_dec_ref(x_118); + x_174 = lean_box(0); } -if (lean_is_scalar(x_156)) { - x_157 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_174)) { + x_175 = lean_alloc_ctor(1, 2, 0); } else { - x_157 = x_156; + x_175 = x_174; } -lean_ctor_set(x_157, 0, x_154); -lean_ctor_set(x_157, 1, x_155); -return x_157; +lean_ctor_set(x_175, 0, x_172); +lean_ctor_set(x_175, 1, x_173); +return x_175; } } else { -lean_object* x_158; -lean_dec(x_103); -x_158 = lean_box(0); -x_5 = x_104; -x_6 = x_158; +lean_object* x_176; +lean_dec(x_115); +x_176 = lean_box(0); +x_5 = x_116; +x_6 = x_176; x_7 = lean_box(0); goto _start; } @@ -4613,7 +4501,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { if (lean_obj_tag(x_4) == 0) @@ -4652,7 +4540,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_1); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_17, x_1, x_1, x_19, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_1, x_2, x_17, x_1, x_1, x_19, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; @@ -4700,7 +4588,7 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { if (lean_obj_tag(x_4) == 0) @@ -4739,7 +4627,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc_n(x_17, 2); -x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(x_17, x_19, x_17, x_17, x_20, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_17, x_19, x_17, x_17, x_20, lean_box(0), x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_17); if (lean_obj_tag(x_21) == 0) { @@ -4803,7 +4691,7 @@ lean_dec(x_10); x_13 = lean_box(0); x_14 = lean_box(0); lean_inc(x_11); -x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(x_11, x_13, x_11, x_11, x_14, lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_15 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_11, x_13, x_11, x_11, x_14, lean_box(0), x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); lean_dec(x_11); if (lean_obj_tag(x_15) == 0) { @@ -4884,43 +4772,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; -x_13 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_13 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -4929,32 +4785,32 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { lean_object* x_17; -x_17 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_17 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); return x_17; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_3); lean_dec(x_2); return x_16; } } -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -5934,7 +5790,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_20 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; lean_ctor_set(x_14, 0, x_20); return x_14; } @@ -5952,7 +5808,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; x_23 = lean_ctor_get(x_21, 0); lean_dec(x_23); -x_24 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_24 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; lean_ctor_set(x_21, 0, x_24); return x_21; } @@ -5962,7 +5818,7 @@ lean_object* x_25; lean_object* x_26; lean_object* x_27; x_25 = lean_ctor_get(x_21, 1); lean_inc(x_25); lean_dec(x_21); -x_26 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_26 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_25); @@ -6016,7 +5872,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_35 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_35 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; x_36 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_32); @@ -6039,7 +5895,7 @@ if (lean_is_exclusive(x_37)) { lean_dec_ref(x_37); x_39 = lean_box(0); } -x_40 = l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1; +x_40 = l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5; if (lean_is_scalar(x_39)) { x_41 = lean_alloc_ctor(0, 2, 0); } else { @@ -6680,61 +6536,58 @@ l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParent lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__1); l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2(); lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__2___closed__2); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__1); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__2); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__3); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__4); +l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___lambda__1___closed__5); l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1(); lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__1); -l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__2); -l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__3); -l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__4); -l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___spec__3___closed__5); l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__1); l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__2); l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkParents___closed__3); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__1); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__2); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__3); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__4); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__5); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__6); -l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__1___closed__7); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__1(); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2(); -lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__2); -l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3 = _init_l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3(); -lean_mark_persistent(l_Lean_addTrace___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__2___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___lambda__1___closed__4); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__2); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__3); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__4); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__5); -l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__3___closed__6); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__1); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__2); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__3); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__4); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__5); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__6); +l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkPtrEqImpliesStructEq___spec__2___closed__7); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___lambda__1___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__1); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__2); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__3); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__4); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__5); +l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6 = _init_l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Meta_Tactic_Grind_Inv_0__Lean_Meta_Grind_checkProofs___spec__1___closed__6); l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_checkInvariants___lambda__1___closed__1); l_Lean_Meta_Grind_checkInvariants___closed__1 = _init_l_Lean_Meta_Grind_checkInvariants___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c index 3b38da4bdebb..b1468223a67f 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Main.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Main -// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Util Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.PropagatorAttr Lean.Meta.Tactic.Grind.Proj Lean.Meta.Tactic.Grind.ForallProp Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.EMatch +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Util Lean.Meta.Tactic.Grind.RevertAll Lean.Meta.Tactic.Grind.PropagatorAttr Lean.Meta.Tactic.Grind.Proj Lean.Meta.Tactic.Grind.ForallProp Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Inv Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.EMatch Lean.Meta.Tactic.Grind.Split Lean.Meta.Tactic.Grind.SimpUtil #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,34 +13,40 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_Meta_Grind_ematchStar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2; +lean_object* l_Lean_Meta_Grind_splitNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); +lean_object* l_Lean_Meta_Grind_getSimprocs___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__1(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_MVarId_revertAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getSimpContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__2; lean_object* l_Std_Queue_empty(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); @@ -52,36 +58,35 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; lean_object* l_Lean_Meta_Grind_Goal_checkInvariants(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_st_ref_take(lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); -lean_object* l_Lean_Meta_Simp_SimprocExtension_getSimprocs(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__3; static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_propagateProjEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_getTrueExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_preprocessAndProbe___closed__1; +lean_object* l_Lean_Meta_appendTagSuffix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_mkMethods___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; lean_object* l_ShareCommon_mkStateImpl(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Simp_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_builtinPropagatorsRef; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; +lean_object* l_Lean_Meta_Grind_GrindTactic_iterate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__5; lean_object* l_Lean_MVarId_clearAuxDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__7; @@ -89,36 +94,35 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Gr static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg(lean_object*); -extern lean_object* l_Lean_Meta_Grind_grindNormSimprocExt; -static lean_object* l_Lean_Meta_Grind_main___closed__1; -lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_GrindTactic_andThen(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__6; +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_ensureNoMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_intros(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1; -static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__19; -static lean_object* l_Lean_Meta_Grind_main___lambda__2___closed__6; +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3; +lean_object* l_Lean_Meta_Grind_propagateImpliesDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); -extern lean_object* l_Lean_Meta_Grind_grindNormExt; uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(lean_object*, lean_object*); lean_object* l_Lean_MVarId_transformTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_resetCache___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__17; lean_object* l_Lean_Meta_Grind_mkENodeCore(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3; lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_all(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); @@ -126,23 +130,23 @@ lean_object* lean_array_mk(lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__4; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Meta_Grind_all___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_ensureProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1; lean_object* lean_array_uget(lean_object*, size_t); -lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_unfoldReducible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; -lean_object* l_Lean_Meta_Grind_propagateForallProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_propagateForallPropUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); extern lean_object* l_Lean_ShareCommon_objectFactory; lean_object* lean_state_sharecommon(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_mkMethods___rarg___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; static lean_object* l_Lean_Meta_Grind_GrindM_run___rarg___closed__2; +lean_object* l_Lean_Meta_Grind_applyToAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_ppGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_betaReduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Match_instInhabitedMatchEqnsExtState___spec__1; size_t lean_usize_land(size_t, size_t); lean_object* l_Lean_Meta_Grind_getEMatchTheorems___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(lean_object* x_1, lean_object* x_2) { @@ -177,7 +181,7 @@ return x_9; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; @@ -190,7 +194,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_12 = l_Lean_Meta_Grind_propagateForallProp(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_Meta_Grind_propagateForallPropUp(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_12) == 0) { uint8_t x_13; @@ -560,49 +564,60 @@ return x_105; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_Expr_getAppFn(x_2); -if (lean_obj_tag(x_12) == 4) +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_12 = l_Lean_Meta_Grind_propagateImpliesDown(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = lean_array_get_size(x_16); -x_19 = l_Lean_Name_hash___override(x_14); -x_20 = 32; -x_21 = lean_uint64_shift_right(x_19, x_20); -x_22 = lean_uint64_xor(x_19, x_21); -x_23 = 16; -x_24 = lean_uint64_shift_right(x_22, x_23); -x_25 = lean_uint64_xor(x_22, x_24); -x_26 = lean_uint64_to_usize(x_25); -x_27 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_28 = 1; -x_29 = lean_usize_sub(x_27, x_28); -x_30 = lean_usize_land(x_26, x_29); -x_31 = lean_array_uget(x_16, x_30); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 1); +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = l_Lean_Expr_getAppFn(x_2); +if (lean_obj_tag(x_16) == 4) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; size_t x_30; size_t x_31; size_t x_32; lean_object* x_33; lean_object* x_34; +x_17 = lean_ctor_get(x_1, 1); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); lean_dec(x_16); -x_32 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_14, x_31); -lean_dec(x_31); -lean_dec(x_14); -if (lean_obj_tag(x_32) == 0) +x_19 = lean_ctor_get(x_17, 1); +x_20 = lean_array_get_size(x_19); +x_21 = l_Lean_Name_hash___override(x_18); +x_22 = 32; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = 16; +x_26 = lean_uint64_shift_right(x_24, x_25); +x_27 = lean_uint64_xor(x_24, x_26); +x_28 = lean_uint64_to_usize(x_27); +x_29 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_30 = 1; +x_31 = lean_usize_sub(x_29, x_30); +x_32 = lean_usize_land(x_28, x_31); +x_33 = lean_array_uget(x_19, x_32); +x_34 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_18, x_33); +lean_dec(x_33); +lean_dec(x_18); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_33; +lean_object* x_35; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -612,50 +627,75 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_33 = lean_box(0); -lean_ctor_set(x_13, 1, x_11); -lean_ctor_set(x_13, 0, x_33); -return x_13; +x_35 = lean_box(0); +lean_ctor_set(x_12, 0, x_35); +return x_12; } else { -lean_object* x_34; lean_object* x_35; -lean_free_object(x_13); -x_34 = lean_ctor_get(x_32, 0); -lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_apply_10(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_35; +lean_object* x_36; lean_object* x_37; +lean_free_object(x_12); +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_apply_10(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_37; } } else { -lean_object* x_36; lean_object* x_37; uint64_t x_38; uint64_t x_39; uint64_t x_40; uint64_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; size_t x_45; size_t x_46; size_t x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; -x_36 = lean_ctor_get(x_13, 1); -lean_inc(x_36); -lean_dec(x_13); -x_37 = lean_array_get_size(x_36); -x_38 = l_Lean_Name_hash___override(x_14); -x_39 = 32; -x_40 = lean_uint64_shift_right(x_38, x_39); -x_41 = lean_uint64_xor(x_38, x_40); -x_42 = 16; -x_43 = lean_uint64_shift_right(x_41, x_42); -x_44 = lean_uint64_xor(x_41, x_43); -x_45 = lean_uint64_to_usize(x_44); -x_46 = lean_usize_of_nat(x_37); -lean_dec(x_37); -x_47 = 1; -x_48 = lean_usize_sub(x_46, x_47); -x_49 = lean_usize_land(x_45, x_48); -x_50 = lean_array_uget(x_36, x_49); -lean_dec(x_36); -x_51 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_14, x_50); -lean_dec(x_50); -lean_dec(x_14); -if (lean_obj_tag(x_51) == 0) +lean_object* x_38; +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_38 = lean_box(0); +lean_ctor_set(x_12, 0, x_38); +return x_12; +} +} +else { -lean_object* x_52; lean_object* x_53; +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_12, 1); +lean_inc(x_39); +lean_dec(x_12); +x_40 = l_Lean_Expr_getAppFn(x_2); +if (lean_obj_tag(x_40) == 4) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; size_t x_52; size_t x_53; size_t x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; +x_41 = lean_ctor_get(x_1, 1); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_ctor_get(x_41, 1); +x_44 = lean_array_get_size(x_43); +x_45 = l_Lean_Name_hash___override(x_42); +x_46 = 32; +x_47 = lean_uint64_shift_right(x_45, x_46); +x_48 = lean_uint64_xor(x_45, x_47); +x_49 = 16; +x_50 = lean_uint64_shift_right(x_48, x_49); +x_51 = lean_uint64_xor(x_48, x_50); +x_52 = lean_uint64_to_usize(x_51); +x_53 = lean_usize_of_nat(x_44); +lean_dec(x_44); +x_54 = 1; +x_55 = lean_usize_sub(x_53, x_54); +x_56 = lean_usize_land(x_52, x_55); +x_57 = lean_array_uget(x_43, x_56); +x_58 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_42, x_57); +lean_dec(x_57); +lean_dec(x_42); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -665,27 +705,46 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_52 = lean_box(0); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_11); -return x_53; +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_39); +return x_60; } else { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_51, 0); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_apply_10(x_54, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_55; +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_58, 0); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_apply_10(x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_39); +return x_62; +} +} +else +{ +lean_object* x_63; lean_object* x_64; +lean_dec(x_40); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_63 = lean_box(0); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_39); +return x_64; } } } else { -lean_object* x_56; lean_object* x_57; -lean_dec(x_12); +uint8_t x_65; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -695,16 +754,28 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_56 = lean_box(0); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_11); -return x_57; +x_65 = !lean_is_exclusive(x_12); +if (x_65 == 0) +{ +return x_12; +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_12, 0); +x_67 = lean_ctor_get(x_12, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_12); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; +} } } } -static lean_object* _init_l_Lean_Meta_Grind_mkMethods___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_Grind_mkMethods___closed__1() { _start: { lean_object* x_1; @@ -712,86 +783,89 @@ x_1 = l_Lean_Meta_Grind_builtinPropagatorsRef; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_Meta_Grind_mkMethods___rarg___closed__1; -x_3 = lean_st_ref_get(x_2, x_1); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = l_Lean_Meta_Grind_mkMethods___closed__1; +x_6 = lean_st_ref_get(x_5, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -x_6 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed), 11, 1); -lean_closure_set(x_6, 0, x_5); -x_7 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__2), 11, 1); -lean_closure_set(x_7, 0, x_5); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_6); -lean_ctor_set(x_8, 1, x_7); -lean_ctor_set(x_3, 0, x_8); -return x_3; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___lambda__1___boxed), 11, 1); +lean_closure_set(x_9, 0, x_8); +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___lambda__2___boxed), 11, 1); +lean_closure_set(x_10, 0, x_8); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set(x_11, 2, x_1); +lean_ctor_set(x_6, 0, x_11); +return x_6; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = lean_ctor_get(x_3, 0); -x_10 = lean_ctor_get(x_3, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_3); -lean_inc(x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed), 11, 1); -lean_closure_set(x_11, 0, x_9); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg___lambda__2), 11, 1); -lean_closure_set(x_12, 0, x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_10); -return x_14; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_12 = lean_ctor_get(x_6, 0); +x_13 = lean_ctor_get(x_6, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_6); +lean_inc(x_12); +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___lambda__1___boxed), 11, 1); +lean_closure_set(x_14, 0, x_12); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___lambda__2___boxed), 11, 1); +lean_closure_set(x_15, 0, x_12); +x_16 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_16, 2, x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_13); +return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_mkMethods___rarg), 1, 0); +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_3; -x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_mkMethods___spec__1(x_1, x_2); -lean_dec(x_2); +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_mkMethods___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_1); -return x_3; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_Meta_Grind_mkMethods___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_Meta_Grind_mkMethods___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_1); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkMethods___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_3; -x_3 = l_Lean_Meta_Grind_mkMethods(x_1, x_2); +lean_object* x_5; +x_5 = l_Lean_Meta_Grind_mkMethods(x_1, x_2, x_3, x_4); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_3; +return x_5; } } static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__1() { @@ -874,75 +948,25 @@ static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__9() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_Grind_grindNormExt; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_grindNormSimprocExt; -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Simp_defaultMaxSteps; -x_2 = lean_unsigned_to_nat(2u); -x_3 = 0; -x_4 = 1; -x_5 = 0; -x_6 = lean_alloc_ctor(0, 2, 19); -lean_ctor_set(x_6, 0, x_1); -lean_ctor_set(x_6, 1, x_2); -lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 1, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 2, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 3, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 4, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 5, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 6, x_5); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 7, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 8, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 9, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 10, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 11, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 12, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 13, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 14, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 15, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 16, x_3); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 17, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 18, x_4); -return x_6; -} -} -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__12() { -_start: -{ -lean_object* x_1; x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__13() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__14() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; x_2 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -950,11 +974,11 @@ lean_ctor_set(x_2, 2, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__15() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -962,7 +986,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__16() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -971,23 +995,23 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__18() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__15() { _start: { size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 5; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__17; -x_3 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; +x_3 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; x_4 = lean_unsigned_to_nat(0u); x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); lean_ctor_set(x_5, 0, x_2); @@ -998,12 +1022,12 @@ lean_ctor_set_usize(x_5, 4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__19() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; x_3 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_1); @@ -1012,23 +1036,22 @@ lean_ctor_set(x_3, 3, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__20() { +static lean_object* _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; -x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__19; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__12; +x_2 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__16; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_9 = lean_box(0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_10 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__5; x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); @@ -1042,358 +1065,140 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__9; -x_19 = l_Lean_Meta_SimpExtension_getTheorems(x_18, x_6, x_7, x_8); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; -x_23 = l_Lean_Meta_Simp_SimprocExtension_getSimprocs(x_22, x_6, x_7, x_21); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +x_18 = l_Lean_Meta_Grind_getSimprocs___rarg(x_7, x_8, x_9); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_ctor_get(x_23, 1); -lean_ctor_set_tag(x_23, 1); -lean_ctor_set(x_23, 1, x_9); -x_26 = lean_array_mk(x_23); -x_27 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_29 = lean_ctor_get(x_27, 0); -x_30 = lean_ctor_get(x_27, 1); -lean_ctor_set_tag(x_27, 1); -lean_ctor_set(x_27, 1, x_9); -lean_ctor_set(x_27, 0, x_20); -x_31 = lean_array_mk(x_27); -x_32 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_33 = l_Lean_Meta_Simp_mkContext(x_32, x_31, x_29, x_4, x_5, x_6, x_7, x_30); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Meta_Grind_getSimpContext(x_5, x_6, x_7, x_8, x_20); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Meta_Grind_mkMethods(x_4, x_7, x_8, x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_27, 0, x_22); +lean_ctor_set(x_27, 1, x_19); +lean_ctor_set(x_27, 2, x_2); +lean_ctor_set(x_27, 3, x_3); +x_28 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; +x_29 = lean_unsigned_to_nat(1u); +x_30 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; +x_31 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__17; +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_33, 0, x_28); +lean_ctor_set(x_33, 1, x_17); +lean_ctor_set(x_33, 2, x_29); +lean_ctor_set(x_33, 3, x_30); +lean_ctor_set(x_33, 4, x_31); +lean_ctor_set(x_33, 5, x_16); +lean_ctor_set(x_33, 6, x_11); +lean_ctor_set(x_33, 7, x_32); +x_34 = lean_st_mk_ref(x_33, x_26); +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_Meta_Grind_mkMethods___rarg(x_35); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_39, 0, x_34); -lean_ctor_set(x_39, 1, x_26); -lean_ctor_set(x_39, 2, x_2); -lean_ctor_set(x_39, 3, x_3); -x_40 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_41 = lean_unsigned_to_nat(1u); -x_42 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_43 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_44 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, x_17); -lean_ctor_set(x_44, 2, x_41); -lean_ctor_set(x_44, 3, x_42); -lean_ctor_set(x_44, 4, x_43); -lean_ctor_set(x_44, 5, x_16); -lean_ctor_set(x_44, 6, x_11); -x_45 = lean_st_mk_ref(x_44, x_38); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -lean_inc(x_46); -x_48 = lean_apply_8(x_1, x_37, x_39, x_46, x_4, x_5, x_6, x_7, x_47); -if (lean_obj_tag(x_48) == 0) +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +lean_inc(x_35); +x_37 = lean_apply_8(x_1, x_25, x_27, x_35, x_5, x_6, x_7, x_8, x_36); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = lean_st_ref_get(x_46, x_50); -lean_dec(x_46); -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_st_ref_get(x_35, x_39); +lean_dec(x_35); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_53; -x_53 = lean_ctor_get(x_51, 0); -lean_dec(x_53); -lean_ctor_set(x_51, 0, x_49); -return x_51; +lean_object* x_42; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +lean_ctor_set(x_40, 0, x_38); +return x_40; } else { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); -lean_dec(x_51); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_49); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_38); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } else { -uint8_t x_56; -lean_dec(x_46); -x_56 = !lean_is_exclusive(x_48); -if (x_56 == 0) +uint8_t x_45; +lean_dec(x_35); +x_45 = !lean_is_exclusive(x_37); +if (x_45 == 0) { -return x_48; +return x_37; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_48, 0); -x_58 = lean_ctor_get(x_48, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_48); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_37, 0); +x_47 = lean_ctor_get(x_37, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_37); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_60 = lean_ctor_get(x_27, 0); -x_61 = lean_ctor_get(x_27, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_27); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_20); -lean_ctor_set(x_62, 1, x_9); -x_63 = lean_array_mk(x_62); -x_64 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_65 = l_Lean_Meta_Simp_mkContext(x_64, x_63, x_60, x_4, x_5, x_6, x_7, x_61); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = l_Lean_Meta_Grind_mkMethods___rarg(x_67); -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_71, 0, x_66); -lean_ctor_set(x_71, 1, x_26); -lean_ctor_set(x_71, 2, x_2); -lean_ctor_set(x_71, 3, x_3); -x_72 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_73 = lean_unsigned_to_nat(1u); -x_74 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_75 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_76 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_76, 0, x_72); -lean_ctor_set(x_76, 1, x_17); -lean_ctor_set(x_76, 2, x_73); -lean_ctor_set(x_76, 3, x_74); -lean_ctor_set(x_76, 4, x_75); -lean_ctor_set(x_76, 5, x_16); -lean_ctor_set(x_76, 6, x_11); -x_77 = lean_st_mk_ref(x_76, x_70); -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec(x_77); -lean_inc(x_78); -x_80 = lean_apply_8(x_1, x_69, x_71, x_78, x_4, x_5, x_6, x_7, x_79); -if (lean_obj_tag(x_80) == 0) +uint8_t x_49; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_49 = !lean_is_exclusive(x_18); +if (x_49 == 0) { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -lean_dec(x_80); -x_83 = lean_st_ref_get(x_78, x_82); -lean_dec(x_78); -x_84 = lean_ctor_get(x_83, 1); -lean_inc(x_84); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - x_85 = x_83; -} else { - lean_dec_ref(x_83); - x_85 = lean_box(0); -} -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(0, 2, 0); -} else { - x_86 = x_85; -} -lean_ctor_set(x_86, 0, x_81); -lean_ctor_set(x_86, 1, x_84); -return x_86; +return x_18; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -lean_dec(x_78); -x_87 = lean_ctor_get(x_80, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_80, 1); -lean_inc(x_88); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_89 = x_80; -} else { - lean_dec_ref(x_80); - x_89 = lean_box(0); -} -if (lean_is_scalar(x_89)) { - x_90 = lean_alloc_ctor(1, 2, 0); -} else { - x_90 = x_89; -} -lean_ctor_set(x_90, 0, x_87); -lean_ctor_set(x_90, 1, x_88); -return x_90; -} -} -} -else -{ -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_91 = lean_ctor_get(x_23, 0); -x_92 = lean_ctor_get(x_23, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_23); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_9); -x_94 = lean_array_mk(x_93); -x_95 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_92); -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_98 = x_95; -} else { - lean_dec_ref(x_95); - x_98 = lean_box(0); -} -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(1, 2, 0); -} else { - x_99 = x_98; - lean_ctor_set_tag(x_99, 1); -} -lean_ctor_set(x_99, 0, x_20); -lean_ctor_set(x_99, 1, x_9); -x_100 = lean_array_mk(x_99); -x_101 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__11; -x_102 = l_Lean_Meta_Simp_mkContext(x_101, x_100, x_96, x_4, x_5, x_6, x_7, x_97); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_105 = l_Lean_Meta_Grind_mkMethods___rarg(x_104); -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -lean_dec(x_105); -x_108 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_108, 0, x_103); -lean_ctor_set(x_108, 1, x_94); -lean_ctor_set(x_108, 2, x_2); -lean_ctor_set(x_108, 3, x_3); -x_109 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__14; -x_110 = lean_unsigned_to_nat(1u); -x_111 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_112 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__20; -x_113 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_113, 0, x_109); -lean_ctor_set(x_113, 1, x_17); -lean_ctor_set(x_113, 2, x_110); -lean_ctor_set(x_113, 3, x_111); -lean_ctor_set(x_113, 4, x_112); -lean_ctor_set(x_113, 5, x_16); -lean_ctor_set(x_113, 6, x_11); -x_114 = lean_st_mk_ref(x_113, x_107); -x_115 = lean_ctor_get(x_114, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 1); -lean_inc(x_116); -lean_dec(x_114); -lean_inc(x_115); -x_117 = lean_apply_8(x_1, x_106, x_108, x_115, x_4, x_5, x_6, x_7, x_116); -if (lean_obj_tag(x_117) == 0) -{ -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_118 = lean_ctor_get(x_117, 0); -lean_inc(x_118); -x_119 = lean_ctor_get(x_117, 1); -lean_inc(x_119); -lean_dec(x_117); -x_120 = lean_st_ref_get(x_115, x_119); -lean_dec(x_115); -x_121 = lean_ctor_get(x_120, 1); -lean_inc(x_121); -if (lean_is_exclusive(x_120)) { - lean_ctor_release(x_120, 0); - lean_ctor_release(x_120, 1); - x_122 = x_120; -} else { - lean_dec_ref(x_120); - x_122 = lean_box(0); -} -if (lean_is_scalar(x_122)) { - x_123 = lean_alloc_ctor(0, 2, 0); -} else { - x_123 = x_122; -} -lean_ctor_set(x_123, 0, x_118); -lean_ctor_set(x_123, 1, x_121); -return x_123; -} -else -{ -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_115); -x_124 = lean_ctor_get(x_117, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_117, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_126 = x_117; -} else { - lean_dec_ref(x_117); - x_126 = lean_box(0); -} -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); -} else { - x_127 = x_126; -} -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_18, 0); +x_51 = lean_ctor_get(x_18, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_18); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } @@ -1402,7 +1207,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GrindM_run(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 8, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 9, 0); return x_2; } } @@ -1410,7 +1215,7 @@ static lean_object* _init_l_Lean_PersistentHashMap_empty___at___private_Lean_Met _start: { lean_object* x_1; -x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; +x_1 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; return x_1; } } @@ -1574,7 +1379,7 @@ return x_1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_10 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_4, x_5, x_6, x_7, x_8, x_9); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); @@ -1593,45 +1398,53 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__13; -x_20 = l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1; -x_21 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1; -x_22 = 0; -x_23 = lean_unsigned_to_nat(0u); -x_24 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__18; -x_25 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; -x_26 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2; +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__10; +x_21 = l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1; +x_22 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1; +x_23 = 0; +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Lean_Meta_Grind_GrindM_run___rarg___closed__15; +x_26 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +x_27 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__2; +x_28 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Match_instInhabitedMatchEqnsExtState___spec__1; +x_29 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3; lean_inc(x_1); -x_27 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_27, 0, x_1); -lean_ctor_set(x_27, 1, x_19); -lean_ctor_set(x_27, 2, x_19); -lean_ctor_set(x_27, 3, x_20); -lean_ctor_set(x_27, 4, x_19); -lean_ctor_set(x_27, 5, x_21); -lean_ctor_set(x_27, 6, x_23); -lean_ctor_set(x_27, 7, x_23); -lean_ctor_set(x_27, 8, x_24); -lean_ctor_set(x_27, 9, x_24); -lean_ctor_set(x_27, 10, x_17); -lean_ctor_set(x_27, 11, x_23); -lean_ctor_set(x_27, 12, x_25); -lean_ctor_set(x_27, 13, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*14, x_22); -x_28 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); -lean_closure_set(x_28, 0, x_27); -x_29 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2___boxed), 11, 2); -lean_closure_set(x_29, 0, x_14); -lean_closure_set(x_29, 1, x_11); -x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_30, 0, x_28); -lean_closure_set(x_30, 1, x_29); -x_31 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; -x_32 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_32, 0, x_30); -lean_closure_set(x_32, 1, x_31); -x_33 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -return x_33; +x_30 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_20); +lean_ctor_set(x_30, 2, x_20); +lean_ctor_set(x_30, 3, x_21); +lean_ctor_set(x_30, 4, x_20); +lean_ctor_set(x_30, 5, x_22); +lean_ctor_set(x_30, 6, x_24); +lean_ctor_set(x_30, 7, x_24); +lean_ctor_set(x_30, 8, x_25); +lean_ctor_set(x_30, 9, x_25); +lean_ctor_set(x_30, 10, x_17); +lean_ctor_set(x_30, 11, x_24); +lean_ctor_set(x_30, 12, x_24); +lean_ctor_set(x_30, 13, x_26); +lean_ctor_set(x_30, 14, x_27); +lean_ctor_set(x_30, 15, x_28); +lean_ctor_set(x_30, 16, x_19); +lean_ctor_set(x_30, 17, x_24); +lean_ctor_set(x_30, 18, x_29); +lean_ctor_set_uint8(x_30, sizeof(void*)*19, x_23); +x_31 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); +lean_closure_set(x_31, 0, x_30); +x_32 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__2___boxed), 11, 2); +lean_closure_set(x_32, 0, x_14); +lean_closure_set(x_32, 1, x_11); +x_33 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_33, 0, x_31); +lean_closure_set(x_33, 1, x_32); +x_34 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; +x_35 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_35, 0, x_33); +lean_closure_set(x_35, 1, x_34); +x_36 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_1, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_36; } } LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -1773,7 +1586,7 @@ else lean_object* x_4; uint8_t x_5; x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); -x_5 = lean_ctor_get_uint8(x_4, sizeof(void*)*14); +x_5 = lean_ctor_get_uint8(x_4, sizeof(void*)*19); if (x_5 == 0) { uint8_t x_6; @@ -1836,6 +1649,24 @@ x_1 = lean_alloc_closure((void*)(l_Lean_MVarId_betaReduce___lambda__1___boxed), return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -1908,12 +1739,21 @@ lean_inc(x_5); x_25 = l_Lean_MVarId_transformTarget(x_22, x_24, x_5, x_6, x_7, x_8, x_23); if (lean_obj_tag(x_25) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); +x_28 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4; +lean_inc(x_26); +x_29 = l_Lean_Meta_appendTagSuffix(x_26, x_28, x_5, x_6, x_7, x_8, x_27); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -1921,16 +1761,16 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_28 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_27); -if (lean_obj_tag(x_28) == 0) +x_31 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_unsigned_to_nat(0u); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_unsigned_to_nat(0u); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -1938,72 +1778,102 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_32 = l_Lean_Meta_Grind_intros(x_29, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -lean_inc(x_33); -x_35 = l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__1(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_34); +x_35 = l_Lean_Meta_Grind_intros(x_34, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_33); if (lean_obj_tag(x_35) == 0) { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +lean_inc(x_36); +x_38 = l_List_forM___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__1(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_35, 0); -lean_dec(x_37); -x_38 = lean_box(0); -x_39 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_33, x_38); -lean_ctor_set(x_35, 0, x_39); -return x_35; +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_38, 0); +lean_dec(x_40); +x_41 = lean_box(0); +x_42 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_36, x_41); +lean_ctor_set(x_38, 0, x_42); +return x_38; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_ctor_get(x_35, 1); -lean_inc(x_40); -lean_dec(x_35); -x_41 = lean_box(0); -x_42 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_33, x_41); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_38, 1); +lean_inc(x_43); +lean_dec(x_38); +x_44 = lean_box(0); +x_45 = l_List_filterTR_loop___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___spec__2(x_36, x_44); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_43); +return x_46; } } else { -uint8_t x_44; -lean_dec(x_33); -x_44 = !lean_is_exclusive(x_35); -if (x_44 == 0) +uint8_t x_47; +lean_dec(x_36); +x_47 = !lean_is_exclusive(x_38); +if (x_47 == 0) +{ +return x_38; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_38, 0); +x_49 = lean_ctor_get(x_38, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_38); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +uint8_t x_51; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_51 = !lean_is_exclusive(x_35); +if (x_51 == 0) { return x_35; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_35, 0); -x_46 = lean_ctor_get(x_35, 1); -lean_inc(x_46); -lean_inc(x_45); +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_35, 0); +x_53 = lean_ctor_get(x_35, 1); +lean_inc(x_53); +lean_inc(x_52); lean_dec(x_35); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } } else { -uint8_t x_48; +uint8_t x_55; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2011,29 +1881,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_48 = !lean_is_exclusive(x_32); -if (x_48 == 0) +x_55 = !lean_is_exclusive(x_31); +if (x_55 == 0) { -return x_32; +return x_31; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_32, 0); -x_50 = lean_ctor_get(x_32, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_32); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_31, 0); +x_57 = lean_ctor_get(x_31, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_31); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; } } } else { -uint8_t x_52; +uint8_t x_59; +lean_dec(x_26); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2041,29 +1912,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_52 = !lean_is_exclusive(x_28); -if (x_52 == 0) +x_59 = !lean_is_exclusive(x_29); +if (x_59 == 0) { -return x_28; +return x_29; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_28, 0); -x_54 = lean_ctor_get(x_28, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_28); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_29, 0); +x_61 = lean_ctor_get(x_29, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_29); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; } } } else { -uint8_t x_56; +uint8_t x_63; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2071,29 +1942,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_56 = !lean_is_exclusive(x_25); -if (x_56 == 0) +x_63 = !lean_is_exclusive(x_25); +if (x_63 == 0) { return x_25; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_25, 0); -x_58 = lean_ctor_get(x_25, 1); -lean_inc(x_58); -lean_inc(x_57); +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_25, 0); +x_65 = lean_ctor_get(x_25, 1); +lean_inc(x_65); +lean_inc(x_64); lean_dec(x_25); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } else { -uint8_t x_60; +uint8_t x_67; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2101,29 +1972,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_60 = !lean_is_exclusive(x_21); -if (x_60 == 0) +x_67 = !lean_is_exclusive(x_21); +if (x_67 == 0) { return x_21; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_21, 0); -x_62 = lean_ctor_get(x_21, 1); -lean_inc(x_62); -lean_inc(x_61); +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_21, 0); +x_69 = lean_ctor_get(x_21, 1); +lean_inc(x_69); +lean_inc(x_68); lean_dec(x_21); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -return x_63; +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; } } } else { -uint8_t x_64; +uint8_t x_71; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2131,29 +2002,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_64 = !lean_is_exclusive(x_17); -if (x_64 == 0) +x_71 = !lean_is_exclusive(x_17); +if (x_71 == 0) { return x_17; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_17, 0); -x_66 = lean_ctor_get(x_17, 1); -lean_inc(x_66); -lean_inc(x_65); +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_17, 0); +x_73 = lean_ctor_get(x_17, 1); +lean_inc(x_73); +lean_inc(x_72); lean_dec(x_17); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -return x_67; +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; } } } else { -uint8_t x_68; +uint8_t x_75; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2161,29 +2032,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_68 = !lean_is_exclusive(x_14); -if (x_68 == 0) +x_75 = !lean_is_exclusive(x_14); +if (x_75 == 0) { return x_14; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_14, 0); -x_70 = lean_ctor_get(x_14, 1); -lean_inc(x_70); -lean_inc(x_69); +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_14, 0); +x_77 = lean_ctor_get(x_14, 1); +lean_inc(x_77); +lean_inc(x_76); lean_dec(x_14); -x_71 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -return x_71; +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } } else { -uint8_t x_72; +uint8_t x_79; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2192,29 +2063,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_72 = !lean_is_exclusive(x_12); -if (x_72 == 0) +x_79 = !lean_is_exclusive(x_12); +if (x_79 == 0) { return x_12; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_12, 0); -x_74 = lean_ctor_get(x_12, 1); -lean_inc(x_74); -lean_inc(x_73); +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_12, 0); +x_81 = lean_ctor_get(x_12, 1); +lean_inc(x_81); +lean_inc(x_80); lean_dec(x_12); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } } else { -uint8_t x_76; +uint8_t x_83; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2223,23 +2094,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_76 = !lean_is_exclusive(x_10); -if (x_76 == 0) +x_83 = !lean_is_exclusive(x_10); +if (x_83 == 0) { return x_10; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_10, 0); -x_78 = lean_ctor_get(x_10, 1); -lean_inc(x_78); -lean_inc(x_77); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_10, 0); +x_85 = lean_ctor_get(x_10, 1); +lean_inc(x_85); +lean_inc(x_84); lean_dec(x_10); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } } } @@ -2338,420 +2209,116 @@ x_12 = l_List_foldlM___at_Lean_Meta_Grind_all___spec__1(x_2, x_11, x_1, x_3, x_4 return x_12; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchStar), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_splitNext), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2() { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; -x_11 = l_Lean_Meta_Grind_all(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_ematchStar), 9, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3() { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindTactic_andThen), 11, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); return x_3; } -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_ctor_get(x_5, 0); -lean_inc(x_7); -lean_dec(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); -{ -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; -} } -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4() { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_box(0); -x_12 = l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__1(x_1, x_11); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_10); -return x_13; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindTactic_iterate), 10, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); -return x_1; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2; +x_11 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4; +x_12 = l_Lean_Meta_Grind_GrindTactic_andThen(x_10, x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("debug", 5, 5); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1), 9, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__3() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("final", 5, 5); -return x_1; +lean_object* x_10; lean_object* x_11; +x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; +x_11 = l_Lean_Meta_Grind_applyToAll(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__4() { +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_main___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_main___lambda__2___closed__2; -x_3 = l_Lean_Meta_Grind_main___lambda__2___closed__3; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__5() { -_start: +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_6, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_ctor_get(x_13, 7); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_10, 0, x_16); +return x_10; } -static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_main___lambda__2___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_10, 0); +x_18 = lean_ctor_get(x_10, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_10); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_19, 7); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_20, x_1); +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +return x_23; +} } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_11 = l_Lean_Meta_Grind_all(x_1, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_main___lambda__2___closed__4; -x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); -x_19 = lean_box(0); -x_20 = l_Lean_Meta_Grind_main___lambda__1(x_12, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_20; -} -else -{ -uint8_t x_21; -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 1); -x_23 = lean_ctor_get(x_15, 0); -lean_dec(x_23); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_12); -x_24 = l_Lean_Meta_Grind_ppGoals(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_22); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_MessageData_ofFormat(x_25); -x_28 = l_Lean_Meta_Grind_main___lambda__2___closed__6; -lean_ctor_set_tag(x_15, 7); -lean_ctor_set(x_15, 1, x_27); -lean_ctor_set(x_15, 0, x_28); -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_15); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_14, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_26); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Meta_Grind_main___lambda__1(x_12, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_31); -return x_33; -} -else -{ -uint8_t x_34; -lean_free_object(x_15); -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_34 = !lean_is_exclusive(x_24); -if (x_34 == 0) -{ -return x_24; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_24, 0); -x_36 = lean_ctor_get(x_24, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_24); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} -} -} -else -{ -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_15, 1); -lean_inc(x_38); -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_12); -x_39 = l_Lean_Meta_Grind_ppGoals(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_38); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = l_Lean_MessageData_ofFormat(x_40); -x_43 = l_Lean_Meta_Grind_main___lambda__2___closed__6; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_43); -x_46 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_14, x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_41); -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); -lean_dec(x_46); -x_49 = l_Lean_Meta_Grind_main___lambda__1(x_12, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_47); -return x_49; -} -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_50 = lean_ctor_get(x_39, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_39, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_52 = x_39; -} else { - lean_dec_ref(x_39); - x_52 = lean_box(0); -} -if (lean_is_scalar(x_52)) { - x_53 = lean_alloc_ctor(1, 2, 0); -} else { - x_53 = x_52; -} -lean_ctor_set(x_53, 0, x_50); -lean_ctor_set(x_53, 1, x_51); -return x_53; -} -} -} -} -else -{ -uint8_t x_54; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_54 = !lean_is_exclusive(x_11); -if (x_54 == 0) -{ -return x_11; -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_11, 0); -x_56 = lean_ctor_get(x_11, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_11); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; -} -} -} -} -static lean_object* _init_l_Lean_Meta_Grind_main___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_main___lambda__2), 9, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore), 9, 1); -lean_closure_set(x_9, 0, x_1); -x_10 = l_Lean_Meta_Grind_main___closed__1; -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Meta_Grind_GrindM_run___rarg(x_11, x_3, x_2, x_4, x_5, x_6, x_7, x_8); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_main___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; @@ -2855,394 +2422,577 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -if (lean_obj_tag(x_2) == 0) -{ lean_object* x_11; lean_object* x_12; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_11 = lean_box(0); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_1); x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); return x_12; } -else +} +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_ctor_get(x_2, 0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_11); +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -x_14 = lean_ctor_get(x_2, 1); -lean_inc(x_14); -lean_dec(x_2); -x_15 = lean_ctor_get(x_13, 0); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); lean_inc(x_15); -x_16 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); -lean_closure_set(x_16, 0, x_13); -lean_inc(x_1); -x_17 = lean_alloc_closure((void*)(l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1___lambda__1), 10, 1); -lean_closure_set(x_17, 0, x_1); -x_18 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_18, 0, x_16); -lean_closure_set(x_18, 1, x_17); -x_19 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; -x_20 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_20, 0, x_18); -lean_closure_set(x_20, 1, x_19); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_12, 0, x_20); +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___lambda__1___boxed), 9, 1); +lean_closure_set(x_13, 0, x_1); +x_14 = lean_alloc_closure((void*)(l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__1), 10, 1); +lean_closure_set(x_14, 0, x_2); +x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_15, 0, x_13); +lean_closure_set(x_15, 1, x_14); +x_16 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__3; +x_17 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_17, 0, x_15); +lean_closure_set(x_17, 1, x_16); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_21 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_15, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_21) == 0) +x_18 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_12, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_22; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_2 = x_14; -x_10 = x_22; -goto _start; -} -else +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get_uint8(x_19, sizeof(void*)*19); +if (x_20 == 0) { -uint8_t x_24; -lean_dec(x_14); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_box(0); +x_23 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3(x_19, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_21); -if (x_24 == 0) -{ -return x_21; +return x_23; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_21, 0); -x_26 = lean_ctor_get(x_21, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_21); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +uint8_t x_24; +lean_dec(x_19); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_18, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_18, 0, x_26); +return x_18; } +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_18, 1); +lean_inc(x_27); +lean_dec(x_18); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } } } +else +{ +uint8_t x_30; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_30 = !lean_is_exclusive(x_18); +if (x_30 == 0) +{ +return x_18; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_18, 0); +x_32 = lean_ctor_get(x_18, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_18); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} } -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +} +} +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_st_ref_get(x_3, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); -x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_12 = l_List_reverse___rarg(x_3); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +x_15 = lean_ctor_get_uint8(x_14, sizeof(void*)*19); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_2); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_2, 1); +x_18 = lean_ctor_get(x_2, 0); +lean_dec(x_18); +x_19 = lean_box(0); lean_inc(x_10); -lean_dec(x_8); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_11 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_11) == 0) +lean_inc(x_1); +x_20 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(x_14, x_1, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_resetCache___rarg(x_3, x_4, x_5, x_13); -lean_dec(x_5); -lean_dec(x_4); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_st_ref_take(x_3, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) +lean_object* x_21; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_17, 0); +lean_object* x_22; +lean_free_object(x_2); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); lean_dec(x_20); -lean_ctor_set(x_17, 0, x_10); -x_21 = lean_st_ref_set(x_3, x_17, x_18); -lean_dec(x_3); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_12); -return x_21; +x_2 = x_17; +x_11 = x_22; +goto _start; } else { lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 1); +x_24 = lean_ctor_get(x_20, 1); lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 0); +lean_inc(x_25); lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_12); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_25); +{ +lean_object* _tmp_1 = x_17; +lean_object* _tmp_2 = x_2; +lean_object* _tmp_10 = x_24; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_11 = _tmp_10; +} +goto _start; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_26 = lean_ctor_get(x_17, 1); -x_27 = lean_ctor_get(x_17, 2); -x_28 = lean_ctor_get(x_17, 3); -x_29 = lean_ctor_get(x_17, 4); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); +uint8_t x_27; +lean_free_object(x_2); lean_dec(x_17); -x_30 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_30, 0, x_10); -lean_ctor_set(x_30, 1, x_26); -lean_ctor_set(x_30, 2, x_27); -lean_ctor_set(x_30, 3, x_28); -lean_ctor_set(x_30, 4, x_29); -x_31 = lean_st_ref_set(x_3, x_30, x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - x_33 = x_31; -} else { - lean_dec_ref(x_31); - x_33 = lean_box(0); +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_20); +if (x_27 == 0) +{ +return x_20; } -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(0, 2, 0); -} else { - x_34 = x_33; +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_20, 0); +x_29 = lean_ctor_get(x_20, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_20); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } -lean_ctor_set(x_34, 0, x_12); -lean_ctor_set(x_34, 1, x_32); -return x_34; } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_35 = lean_ctor_get(x_11, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_11, 1); -lean_inc(x_36); -lean_dec(x_11); -x_37 = l_Lean_Meta_resetCache___rarg(x_3, x_4, x_5, x_36); -lean_dec(x_5); -lean_dec(x_4); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = lean_st_ref_take(x_3, x_38); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = !lean_is_exclusive(x_40); -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_40, 0); -lean_dec(x_43); -lean_ctor_set(x_40, 0, x_10); -x_44 = lean_st_ref_set(x_3, x_40, x_41); -lean_dec(x_3); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_dec(x_2); +x_32 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_33 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(x_14, x_1, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_46; -x_46 = lean_ctor_get(x_44, 0); -lean_dec(x_46); -lean_ctor_set_tag(x_44, 1); -lean_ctor_set(x_44, 0, x_35); -return x_44; +lean_object* x_34; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_2 = x_31; +x_11 = x_35; +goto _start; } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_35); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_ctor_get(x_34, 0); +lean_inc(x_38); +lean_dec(x_34); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_3); +x_2 = x_31; +x_3 = x_39; +x_11 = x_37; +goto _start; } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_49 = lean_ctor_get(x_40, 1); -x_50 = lean_ctor_get(x_40, 2); -x_51 = lean_ctor_get(x_40, 3); -x_52 = lean_ctor_get(x_40, 4); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_40); -x_53 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_53, 0, x_10); -lean_ctor_set(x_53, 1, x_49); -lean_ctor_set(x_53, 2, x_50); -lean_ctor_set(x_53, 3, x_51); -lean_ctor_set(x_53, 4, x_52); -x_54 = lean_st_ref_set(x_3, x_53, x_41); +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_dec(x_31); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - x_56 = x_54; +lean_dec(x_1); +x_41 = lean_ctor_get(x_33, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_33, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + x_43 = x_33; } else { - lean_dec_ref(x_54); - x_56 = lean_box(0); + lean_dec_ref(x_33); + x_43 = lean_box(0); } -if (lean_is_scalar(x_56)) { - x_57 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(1, 2, 0); } else { - x_57 = x_56; - lean_ctor_set_tag(x_57, 1); -} -lean_ctor_set(x_57, 0, x_35); -lean_ctor_set(x_57, 1, x_55); -return x_57; + x_44 = x_43; } +lean_ctor_set(x_44, 0, x_41); +lean_ctor_set(x_44, 1, x_42); +return x_44; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg), 6, 0); -return x_2; +lean_object* x_45; +lean_dec(x_14); +x_45 = lean_ctor_get(x_2, 1); +lean_inc(x_45); +lean_dec(x_2); +x_2 = x_45; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__3(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_12; -x_12 = l_List_forM___at_Lean_Meta_Grind_preprocessAndProbe___spec__1(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_12, 0, x_15); -return x_12; +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_12, 1); -lean_inc(x_16); -lean_dec(x_12); -x_17 = lean_box(0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; } +goto _start; } else { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; +x_12 = l_List_mapTR_loop___at_Lean_Meta_Grind_main___spec__3(x_1, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("final", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__3() { +_start: { -return x_12; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3; +x_2 = l_Lean_Meta_Grind_main___lambda__2___closed__1; +x_3 = l_Lean_Meta_Grind_main___lambda__2___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__4() { +_start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_12, 0); -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_12); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_main___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_main___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = l_Lean_Meta_Grind_main___lambda__2___closed__4; -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_12 = l_Lean_Meta_Grind_applyToAll(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); -x_16 = lean_box(0); -x_17 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -return x_17; +x_15 = lean_box(0); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_16 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2(x_1, x_13, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Meta_Grind_main___lambda__2___closed__3; +x_20 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_18); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_unbox(x_21); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_box(0); +x_25 = l_Lean_Meta_Grind_main___lambda__1(x_17, x_15, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_25; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_20); +if (x_26 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_12, 1); -x_20 = lean_ctor_get(x_12, 0); -lean_dec(x_20); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_20, 1); +x_28 = lean_ctor_get(x_20, 0); +lean_dec(x_28); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -3250,38 +3000,46 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_21 = l_Lean_Meta_Grind_ppGoals(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_21) == 0) +lean_inc(x_17); +x_29 = l_Lean_Meta_Grind_ppGoals(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Lean_MessageData_ofFormat(x_22); -x_25 = l_Lean_Meta_Grind_main___lambda__2___closed__6; -lean_ctor_set_tag(x_12, 7); -lean_ctor_set(x_12, 1, x_24); -lean_ctor_set(x_12, 0, x_25); -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_12); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_11, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -lean_dec(x_28); -return x_30; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Lean_MessageData_ofFormat(x_30); +x_33 = l_Lean_Meta_Grind_main___lambda__2___closed__5; +lean_ctor_set_tag(x_20, 7); +lean_ctor_set(x_20, 1, x_32); +lean_ctor_set(x_20, 0, x_33); +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_20); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_19, x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_Lean_Meta_Grind_main___lambda__1(x_17, x_15, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_37); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_36); +return x_38; } else { -uint8_t x_31; -lean_free_object(x_12); +uint8_t x_39; +lean_free_object(x_20); +lean_dec(x_17); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -3289,34 +3047,32 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_21); -if (x_31 == 0) +x_39 = !lean_is_exclusive(x_29); +if (x_39 == 0) { -return x_21; +return x_29; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_21, 0); -x_33 = lean_ctor_get(x_21, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_21); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_29, 0); +x_41 = lean_ctor_get(x_29, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_29); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } else { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_12, 1); -lean_inc(x_35); -lean_dec(x_12); +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_20, 1); +lean_inc(x_43); +lean_dec(x_20); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -3324,37 +3080,45 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_36 = l_Lean_Meta_Grind_ppGoals(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_35); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_17); +x_44 = l_Lean_Meta_Grind_ppGoals(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = l_Lean_MessageData_ofFormat(x_37); -x_40 = l_Lean_Meta_Grind_main___lambda__2___closed__6; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_11, x_42, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -lean_dec(x_43); -x_46 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); lean_dec(x_44); -return x_46; +x_47 = l_Lean_MessageData_ofFormat(x_45); +x_48 = l_Lean_Meta_Grind_main___lambda__2___closed__5; +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_48); +x_51 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_19, x_50, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_46); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_Meta_Grind_main___lambda__1(x_17, x_15, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_53); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_52); +return x_54; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_17); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -3362,75 +3126,177 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_47 = lean_ctor_get(x_36, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_36, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_49 = x_36; +x_55 = lean_ctor_get(x_44, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_44, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_57 = x_44; } else { - lean_dec_ref(x_36); - x_49 = lean_box(0); + lean_dec_ref(x_44); + x_57 = lean_box(0); } -if (lean_is_scalar(x_49)) { - x_50 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(1, 2, 0); } else { - x_50 = x_49; + x_58 = x_57; } -lean_ctor_set(x_50, 0, x_47); -lean_ctor_set(x_50, 1, x_48); -return x_50; +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_56); +return x_58; +} +} +} +} +else +{ +uint8_t x_59; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_59 = !lean_is_exclusive(x_16); +if (x_59 == 0) +{ +return x_16; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_16, 0); +x_61 = lean_ctor_get(x_16, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_16); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +else +{ +uint8_t x_63; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_12); +if (x_63 == 0) +{ +return x_12; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_12, 0); +x_65 = lean_ctor_get(x_12, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_12); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore), 9, 1); +lean_closure_set(x_10, 0, x_1); +lean_inc(x_4); +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_main___lambda__2), 10, 1); +lean_closure_set(x_11, 0, x_4); +x_12 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_12, 0, x_10); +lean_closure_set(x_12, 1, x_11); +x_13 = l_Lean_Meta_Grind_GrindM_run___rarg(x_12, x_3, x_2, x_4, x_5, x_6, x_7, x_8, x_9); +return x_13; +} } -static lean_object* _init_l_Lean_Meta_Grind_preprocessAndProbe___closed__1() { +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = 0; -x_2 = lean_unsigned_to_nat(100u); -x_3 = lean_unsigned_to_nat(5u); -x_4 = lean_unsigned_to_nat(1000u); -x_5 = lean_alloc_ctor(0, 4, 1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_3); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_uint8(x_5, sizeof(void*)*4, x_1); -return x_5; +lean_object* x_10; +x_10 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_main___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore), 9, 1); -lean_closure_set(x_9, 0, x_1); -x_10 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_preprocessAndProbe___lambda__2), 10, 1); -lean_closure_set(x_10, 0, x_3); -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Meta_Grind_preprocessAndProbe___closed__1; -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_GrindM_run___rarg), 8, 3); -lean_closure_set(x_13, 0, x_11); -lean_closure_set(x_13, 1, x_2); -lean_closure_set(x_13, 2, x_12); -x_14 = l_Lean_Meta_withoutModifyingMCtx___at_Lean_Meta_Grind_preprocessAndProbe___spec__2___rarg(x_13, x_4, x_5, x_6, x_7, x_8); -return x_14; +lean_object* x_11; +x_11 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_preprocessAndProbe___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_List_filterMapM_loop___at_Lean_Meta_Grind_main___spec__2___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_main___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_Meta_Grind_preprocessAndProbe___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_Meta_Grind_main___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); return x_12; } @@ -3445,6 +3311,8 @@ lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object lean_object* initialize_Lean_Meta_Tactic_Grind_Inv(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_EMatch(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Split(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_SimpUtil(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Main(uint8_t builtin, lean_object* w) { lean_object * res; @@ -3480,8 +3348,14 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_EMatch(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Meta_Grind_mkMethods___rarg___closed__1 = _init_l_Lean_Meta_Grind_mkMethods___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_mkMethods___rarg___closed__1); +res = initialize_Lean_Meta_Tactic_Grind_Split(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_SimpUtil(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_mkMethods___closed__1 = _init_l_Lean_Meta_Grind_mkMethods___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_mkMethods___closed__1); l_Lean_Meta_Grind_GrindM_run___rarg___closed__1 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__1); l_Lean_Meta_Grind_GrindM_run___rarg___closed__2 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__2(); @@ -3516,12 +3390,6 @@ l_Lean_Meta_Grind_GrindM_run___rarg___closed__16 = _init_l_Lean_Meta_Grind_Grind lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__16); l_Lean_Meta_Grind_GrindM_run___rarg___closed__17 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__17(); lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__17); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__18 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__18(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__18); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__19 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__19(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__19); -l_Lean_Meta_Grind_GrindM_run___rarg___closed__20 = _init_l_Lean_Meta_Grind_GrindM_run___rarg___closed__20(); -lean_mark_persistent(l_Lean_Meta_Grind_GrindM_run___rarg___closed__20); l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1 = _init_l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1(); lean_mark_persistent(l_Lean_PersistentHashMap_empty___at___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___spec__1); l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_mkGoal___closed__1(); @@ -3534,6 +3402,18 @@ l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1 lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__1); l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__2); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__3); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_initCore___closed__4); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___lambda__1___closed__4); l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Main_0__Lean_Meta_Grind_simple___closed__1); l_Lean_Meta_Grind_main___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__1(); @@ -3546,12 +3426,6 @@ l_Lean_Meta_Grind_main___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_main___ lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__4); l_Lean_Meta_Grind_main___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__5(); lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__5); -l_Lean_Meta_Grind_main___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_main___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Meta_Grind_main___lambda__2___closed__6); -l_Lean_Meta_Grind_main___closed__1 = _init_l_Lean_Meta_Grind_main___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_main___closed__1); -l_Lean_Meta_Grind_preprocessAndProbe___closed__1 = _init_l_Lean_Meta_Grind_preprocessAndProbe___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_preprocessAndProbe___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c index 2f082515ee08..cf887b54feee 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/MarkNestedProofs.c @@ -14,4872 +14,1913 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2; +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofs_unsafe__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); size_t lean_uint64_to_usize(uint64_t); +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isApp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); +static lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2; uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkPtrMap___rarg(lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); uint64_t lean_usize_to_uint64(size_t); +static lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2; lean_object* lean_nat_div(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_outOfBounds___rarg(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3___boxed(lean_object*, lean_object*); +uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2; extern lean_object* l_Lean_levelZero; -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1; extern lean_object* l_Lean_instInhabitedExpr; -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5; -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1; extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl___closed__1; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +uint8_t l_Lean_Expr_isProj(lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(lean_object*); size_t lean_usize_sub(size_t, size_t); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2; -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isForall(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); -static lean_object* _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(uint8_t, uint8_t); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMonadMetaM; -x_2 = l_ReaderT_instMonad___rarg(x_1); -return x_2; -} -} -static lean_object* _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1; -x_2 = l_Lean_instInhabitedExpr; -x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2; -x_9 = lean_panic_fn(x_8, x_1); -x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_2, 1); -x_14 = lean_nat_dec_lt(x_4, x_13); -if (x_14 == 0) -{ -lean_object* x_15; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_3); -lean_ctor_set(x_15, 1, x_12); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; uint8_t x_23; -x_23 = !lean_is_exclusive(x_3); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_ctor_get(x_3, 0); -x_25 = lean_ctor_get(x_3, 1); -x_26 = lean_array_get_size(x_24); -x_27 = lean_nat_dec_lt(x_4, x_26); -lean_dec(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = l_Lean_instInhabitedExpr; -x_29 = l_outOfBounds___rarg(x_28); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_29); -x_30 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_29, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; uint8_t x_35; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = lean_ptr_addr(x_29); -lean_dec(x_29); -x_34 = lean_ptr_addr(x_31); -x_35 = lean_usize_dec_eq(x_33, x_34); -if (x_35 == 0) -{ -lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_25); -x_36 = lean_array_set(x_24, x_4, x_31); -x_37 = 1; -x_38 = lean_box(x_37); -lean_ctor_set(x_3, 1, x_38); -lean_ctor_set(x_3, 0, x_36); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_3); -x_16 = x_39; -x_17 = x_32; -goto block_22; -} -else -{ -lean_object* x_40; -lean_dec(x_31); -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_3); -x_16 = x_40; -x_17 = x_32; -goto block_22; -} -} -else -{ -uint8_t x_41; -lean_dec(x_29); -lean_free_object(x_3); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_41 = !lean_is_exclusive(x_30); -if (x_41 == 0) -{ -return x_30; -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_30, 0); -x_43 = lean_ctor_get(x_30, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_30); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; -} -} -} -else -{ -lean_object* x_45; lean_object* x_46; -x_45 = lean_array_fget(x_24, x_4); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_45); -x_46 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_45, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; uint8_t x_51; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); -lean_dec(x_46); -x_49 = lean_ptr_addr(x_45); -lean_dec(x_45); -x_50 = lean_ptr_addr(x_47); -x_51 = lean_usize_dec_eq(x_49, x_50); -if (x_51 == 0) -{ -lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; -lean_dec(x_25); -x_52 = lean_array_set(x_24, x_4, x_47); -x_53 = 1; -x_54 = lean_box(x_53); -lean_ctor_set(x_3, 1, x_54); -lean_ctor_set(x_3, 0, x_52); -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_3); -x_16 = x_55; -x_17 = x_48; -goto block_22; -} -else -{ -lean_object* x_56; -lean_dec(x_47); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_3); -x_16 = x_56; -x_17 = x_48; -goto block_22; -} -} -else -{ -uint8_t x_57; -lean_dec(x_45); -lean_free_object(x_3); -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_57 = !lean_is_exclusive(x_46); -if (x_57 == 0) -{ -return x_46; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_46, 0); -x_59 = lean_ctor_get(x_46, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_46); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; -} -} -} -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_61 = lean_ctor_get(x_3, 0); -x_62 = lean_ctor_get(x_3, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_3); -x_63 = lean_array_get_size(x_61); -x_64 = lean_nat_dec_lt(x_4, x_63); -lean_dec(x_63); -if (x_64 == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = l_Lean_instInhabitedExpr; -x_66 = l_outOfBounds___rarg(x_65); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_66); -x_67 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_66, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_67) == 0) -{ -lean_object* x_68; lean_object* x_69; size_t x_70; size_t x_71; uint8_t x_72; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -x_70 = lean_ptr_addr(x_66); -lean_dec(x_66); -x_71 = lean_ptr_addr(x_68); -x_72 = lean_usize_dec_eq(x_70, x_71); -if (x_72 == 0) -{ -lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_62); -x_73 = lean_array_set(x_61, x_4, x_68); -x_74 = 1; -x_75 = lean_box(x_74); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_73); -lean_ctor_set(x_76, 1, x_75); -x_77 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_77, 0, x_76); -x_16 = x_77; -x_17 = x_69; -goto block_22; -} -else -{ -lean_object* x_78; lean_object* x_79; -lean_dec(x_68); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_61); -lean_ctor_set(x_78, 1, x_62); -x_79 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_79, 0, x_78); -x_16 = x_79; -x_17 = x_69; -goto block_22; -} -} -else -{ -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_66); -lean_dec(x_62); -lean_dec(x_61); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_80 = lean_ctor_get(x_67, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_67, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_82 = x_67; -} else { - lean_dec_ref(x_67); - x_82 = lean_box(0); -} -if (lean_is_scalar(x_82)) { - x_83 = lean_alloc_ctor(1, 2, 0); -} else { - x_83 = x_82; -} -lean_ctor_set(x_83, 0, x_80); -lean_ctor_set(x_83, 1, x_81); -return x_83; -} -} -else -{ -lean_object* x_84; lean_object* x_85; -x_84 = lean_array_fget(x_61, x_4); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_84); -x_85 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_84, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_85) == 0) -{ -lean_object* x_86; lean_object* x_87; size_t x_88; size_t x_89; uint8_t x_90; -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_85, 1); -lean_inc(x_87); -lean_dec(x_85); -x_88 = lean_ptr_addr(x_84); -lean_dec(x_84); -x_89 = lean_ptr_addr(x_86); -x_90 = lean_usize_dec_eq(x_88, x_89); -if (x_90 == 0) -{ -lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_dec(x_62); -x_91 = lean_array_set(x_61, x_4, x_86); -x_92 = 1; -x_93 = lean_box(x_92); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_91); -lean_ctor_set(x_94, 1, x_93); -x_95 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_95, 0, x_94); -x_16 = x_95; -x_17 = x_87; -goto block_22; -} -else -{ -lean_object* x_96; lean_object* x_97; -lean_dec(x_86); -x_96 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_96, 0, x_61); -lean_ctor_set(x_96, 1, x_62); -x_97 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_97, 0, x_96); -x_16 = x_97; -x_17 = x_87; -goto block_22; -} -} -else -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec(x_84); -lean_dec(x_62); -lean_dec(x_61); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -x_98 = lean_ctor_get(x_85, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_85, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_100 = x_85; -} else { - lean_dec_ref(x_85); - x_100 = lean_box(0); -} -if (lean_is_scalar(x_100)) { - x_101 = lean_alloc_ctor(1, 2, 0); -} else { - x_101 = x_100; -} -lean_ctor_set(x_101, 0, x_98); -lean_ctor_set(x_101, 1, x_99); -return x_101; -} -} -} -block_22: -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_2, 2); -x_20 = lean_nat_add(x_4, x_19); -lean_dec(x_4); -x_3 = x_18; -x_4 = x_20; -x_5 = lean_box(0); -x_6 = lean_box(0); -x_12 = x_17; -goto _start; -} -} -} -} -LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 0; -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_ptr_addr(x_4); -x_7 = lean_ptr_addr(x_1); -x_8 = lean_usize_dec_eq(x_6, x_7); -if (x_8 == 0) -{ -x_2 = x_5; -goto _start; -} -else -{ -uint8_t x_10; -x_10 = 1; -return x_10; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_dec(x_1); -return x_2; -} -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 2); -x_7 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_5); -x_8 = lean_apply_1(x_1, x_5); -x_9 = lean_unbox_uint64(x_8); -lean_dec(x_8); -x_10 = 32; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = 16; -x_14 = lean_uint64_shift_right(x_12, x_13); -x_15 = lean_uint64_xor(x_12, x_14); -x_16 = lean_uint64_to_usize(x_15); -x_17 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_18 = 1; -x_19 = lean_usize_sub(x_17, x_18); -x_20 = lean_usize_land(x_16, x_19); -x_21 = lean_array_uget(x_2, x_20); -lean_ctor_set(x_3, 2, x_21); -x_22 = lean_array_uset(x_2, x_20, x_3); -x_2 = x_22; -x_3 = x_6; -goto _start; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_24 = lean_ctor_get(x_3, 0); -x_25 = lean_ctor_get(x_3, 1); -x_26 = lean_ctor_get(x_3, 2); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_3); -x_27 = lean_array_get_size(x_2); -lean_inc(x_1); -lean_inc(x_24); -x_28 = lean_apply_1(x_1, x_24); -x_29 = lean_unbox_uint64(x_28); -lean_dec(x_28); -x_30 = 32; -x_31 = lean_uint64_shift_right(x_29, x_30); -x_32 = lean_uint64_xor(x_29, x_31); -x_33 = 16; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = lean_uint64_to_usize(x_35); -x_37 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_38 = 1; -x_39 = lean_usize_sub(x_37, x_38); -x_40 = lean_usize_land(x_36, x_39); -x_41 = lean_array_uget(x_2, x_40); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_24); -lean_ctor_set(x_42, 1, x_25); -lean_ctor_set(x_42, 2, x_41); -x_43 = lean_array_uset(x_2, x_40, x_42); -x_2 = x_43; -x_3 = x_26; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -return x_1; -} -else -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 2); -x_6 = lean_array_get_size(x_1); -x_7 = lean_ptr_addr(x_4); -x_8 = lean_usize_to_uint64(x_7); -x_9 = 11; -x_10 = lean_uint64_mix_hash(x_8, x_9); -x_11 = 32; -x_12 = lean_uint64_shift_right(x_10, x_11); -x_13 = lean_uint64_xor(x_10, x_12); -x_14 = 16; -x_15 = lean_uint64_shift_right(x_13, x_14); -x_16 = lean_uint64_xor(x_13, x_15); -x_17 = lean_uint64_to_usize(x_16); -x_18 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_19 = 1; -x_20 = lean_usize_sub(x_18, x_19); -x_21 = lean_usize_land(x_17, x_20); -x_22 = lean_array_uget(x_1, x_21); -lean_ctor_set(x_2, 2, x_22); -x_23 = lean_array_uset(x_1, x_21, x_2); -x_1 = x_23; -x_2 = x_5; -goto _start; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_25 = lean_ctor_get(x_2, 0); -x_26 = lean_ctor_get(x_2, 1); -x_27 = lean_ctor_get(x_2, 2); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_2); -x_28 = lean_array_get_size(x_1); -x_29 = lean_ptr_addr(x_25); -x_30 = lean_usize_to_uint64(x_29); -x_31 = 11; -x_32 = lean_uint64_mix_hash(x_30, x_31); -x_33 = 32; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = 16; -x_37 = lean_uint64_shift_right(x_35, x_36); -x_38 = lean_uint64_xor(x_35, x_37); -x_39 = lean_uint64_to_usize(x_38); -x_40 = lean_usize_of_nat(x_28); -lean_dec(x_28); -x_41 = 1; -x_42 = lean_usize_sub(x_40, x_41); -x_43 = lean_usize_land(x_39, x_42); -x_44 = lean_array_uget(x_1, x_43); -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_25); -lean_ctor_set(x_45, 1, x_26); -lean_ctor_set(x_45, 2, x_44); -x_46 = lean_array_uset(x_1, x_43, x_45); -x_1 = x_46; -x_2 = x_27; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_2); -x_5 = lean_nat_dec_lt(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) -{ -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_6 = lean_array_fget(x_2, x_1); -x_7 = lean_box(0); -x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(x_3, x_6); -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_add(x_1, x_10); -lean_dec(x_1); -x_1 = x_11; -x_2 = x_8; -x_3 = x_9; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(2u); -x_4 = lean_nat_mul(x_2, x_3); -lean_dec(x_2); -x_5 = lean_box(0); -x_6 = lean_mk_array(x_4, x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(x_7, x_1, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_4; -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; -} -else -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 1); -x_8 = lean_ctor_get(x_3, 2); -x_9 = lean_ptr_addr(x_6); -x_10 = lean_ptr_addr(x_1); -x_11 = lean_usize_dec_eq(x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; -x_12 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_2, x_8); -lean_ctor_set(x_3, 2, x_12); -return x_3; -} -else -{ -lean_dec(x_7); -lean_dec(x_6); -lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_1); -return x_3; -} -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; -x_13 = lean_ctor_get(x_3, 0); -x_14 = lean_ctor_get(x_3, 1); -x_15 = lean_ctor_get(x_3, 2); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_3); -x_16 = lean_ptr_addr(x_13); -x_17 = lean_ptr_addr(x_1); -x_18 = lean_usize_dec_eq(x_16, x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_2, x_15); -x_20 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_14); -lean_ctor_set(x_20, 2, x_19); -return x_20; -} -else -{ -lean_object* x_21; -lean_dec(x_14); -lean_dec(x_13); -x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_1); -lean_ctor_set(x_21, 1, x_2); -lean_ctor_set(x_21, 2, x_15); -return x_21; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -switch (lean_obj_tag(x_2)) { -case 0: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_4); -x_11 = lean_array_get_size(x_3); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_11); -lean_ctor_set(x_14, 2, x_13); -x_15 = 0; -x_16 = lean_box(x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_3); -lean_ctor_set(x_17, 1, x_16); -lean_inc(x_5); -x_18 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_14, x_14, x_17, x_12, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_125; uint8_t x_126; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_125 = lean_ctor_get(x_19, 1); -lean_inc(x_125); -x_126 = lean_unbox(x_125); -lean_dec(x_125); -if (x_126 == 0) -{ -lean_dec(x_19); -lean_dec(x_2); -lean_inc(x_1); -x_21 = x_1; -goto block_124; -} -else -{ -lean_object* x_127; lean_object* x_128; -x_127 = lean_ctor_get(x_19, 0); -lean_inc(x_127); -lean_dec(x_19); -x_128 = l_Lean_mkAppN(x_2, x_127); -lean_dec(x_127); -x_21 = x_128; -goto block_124; -} -block_124: -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_st_ref_take(x_5, x_20); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; uint8_t x_45; -x_26 = lean_ctor_get(x_23, 0); -x_27 = lean_ctor_get(x_23, 1); -x_28 = lean_array_get_size(x_27); -x_29 = lean_ptr_addr(x_1); -x_30 = lean_usize_to_uint64(x_29); -x_31 = 11; -x_32 = lean_uint64_mix_hash(x_30, x_31); -x_33 = 32; -x_34 = lean_uint64_shift_right(x_32, x_33); -x_35 = lean_uint64_xor(x_32, x_34); -x_36 = 16; -x_37 = lean_uint64_shift_right(x_35, x_36); -x_38 = lean_uint64_xor(x_35, x_37); -x_39 = lean_uint64_to_usize(x_38); -x_40 = lean_usize_of_nat(x_28); -lean_dec(x_28); -x_41 = 1; -x_42 = lean_usize_sub(x_40, x_41); -x_43 = lean_usize_land(x_39, x_42); -x_44 = lean_array_uget(x_27, x_43); -x_45 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_44); -if (x_45 == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_46 = lean_nat_add(x_26, x_13); -lean_dec(x_26); -lean_inc(x_21); -x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_1); -lean_ctor_set(x_47, 1, x_21); -lean_ctor_set(x_47, 2, x_44); -x_48 = lean_array_uset(x_27, x_43, x_47); -x_49 = lean_unsigned_to_nat(4u); -x_50 = lean_nat_mul(x_46, x_49); -x_51 = lean_unsigned_to_nat(3u); -x_52 = lean_nat_div(x_50, x_51); -lean_dec(x_50); -x_53 = lean_array_get_size(x_48); -x_54 = lean_nat_dec_le(x_52, x_53); -lean_dec(x_53); -lean_dec(x_52); -if (x_54 == 0) -{ -lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_55 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_48); -lean_ctor_set(x_23, 1, x_55); -lean_ctor_set(x_23, 0, x_46); -x_56 = lean_st_ref_set(x_5, x_23, x_24); -lean_dec(x_5); -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) -{ -lean_object* x_58; -x_58 = lean_ctor_get(x_56, 0); -lean_dec(x_58); -lean_ctor_set(x_56, 0, x_21); -return x_56; -} -else -{ -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_21); -lean_ctor_set(x_60, 1, x_59); -return x_60; -} -} -else -{ -lean_object* x_61; uint8_t x_62; -lean_ctor_set(x_23, 1, x_48); -lean_ctor_set(x_23, 0, x_46); -x_61 = lean_st_ref_set(x_5, x_23, x_24); -lean_dec(x_5); -x_62 = !lean_is_exclusive(x_61); -if (x_62 == 0) -{ -lean_object* x_63; -x_63 = lean_ctor_get(x_61, 0); -lean_dec(x_63); -lean_ctor_set(x_61, 0, x_21); -return x_61; -} -else -{ -lean_object* x_64; lean_object* x_65; -x_64 = lean_ctor_get(x_61, 1); -lean_inc(x_64); -lean_dec(x_61); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_21); -lean_ctor_set(x_65, 1, x_64); -return x_65; -} -} -} -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_66 = lean_box(0); -x_67 = lean_array_uset(x_27, x_43, x_66); -lean_inc(x_21); -x_68 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_21, x_44); -x_69 = lean_array_uset(x_67, x_43, x_68); -lean_ctor_set(x_23, 1, x_69); -x_70 = lean_st_ref_set(x_5, x_23, x_24); -lean_dec(x_5); -x_71 = !lean_is_exclusive(x_70); -if (x_71 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_72; -x_72 = lean_ctor_get(x_70, 0); -lean_dec(x_72); -lean_ctor_set(x_70, 0, x_21); -return x_70; +uint8_t x_3; +x_3 = 0; +return x_3; } else { -lean_object* x_73; lean_object* x_74; -x_73 = lean_ctor_get(x_70, 1); -lean_inc(x_73); -lean_dec(x_70); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_21); -lean_ctor_set(x_74, 1, x_73); -return x_74; -} -} -} -else +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ptr_addr(x_4); +x_7 = lean_ptr_addr(x_1); +x_8 = lean_usize_dec_eq(x_6, x_7); +if (x_8 == 0) { -lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; uint64_t x_79; uint64_t x_80; uint64_t x_81; uint64_t x_82; uint64_t x_83; uint64_t x_84; uint64_t x_85; uint64_t x_86; uint64_t x_87; size_t x_88; size_t x_89; size_t x_90; size_t x_91; size_t x_92; lean_object* x_93; uint8_t x_94; -x_75 = lean_ctor_get(x_23, 0); -x_76 = lean_ctor_get(x_23, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_23); -x_77 = lean_array_get_size(x_76); -x_78 = lean_ptr_addr(x_1); -x_79 = lean_usize_to_uint64(x_78); -x_80 = 11; -x_81 = lean_uint64_mix_hash(x_79, x_80); -x_82 = 32; -x_83 = lean_uint64_shift_right(x_81, x_82); -x_84 = lean_uint64_xor(x_81, x_83); -x_85 = 16; -x_86 = lean_uint64_shift_right(x_84, x_85); -x_87 = lean_uint64_xor(x_84, x_86); -x_88 = lean_uint64_to_usize(x_87); -x_89 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_90 = 1; -x_91 = lean_usize_sub(x_89, x_90); -x_92 = lean_usize_land(x_88, x_91); -x_93 = lean_array_uget(x_76, x_92); -x_94 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_93); -if (x_94 == 0) -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_95 = lean_nat_add(x_75, x_13); -lean_dec(x_75); -lean_inc(x_21); -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_1); -lean_ctor_set(x_96, 1, x_21); -lean_ctor_set(x_96, 2, x_93); -x_97 = lean_array_uset(x_76, x_92, x_96); -x_98 = lean_unsigned_to_nat(4u); -x_99 = lean_nat_mul(x_95, x_98); -x_100 = lean_unsigned_to_nat(3u); -x_101 = lean_nat_div(x_99, x_100); -lean_dec(x_99); -x_102 = lean_array_get_size(x_97); -x_103 = lean_nat_dec_le(x_101, x_102); -lean_dec(x_102); -lean_dec(x_101); -if (x_103 == 0) -{ -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_104 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_97); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_95); -lean_ctor_set(x_105, 1, x_104); -x_106 = lean_st_ref_set(x_5, x_105, x_24); -lean_dec(x_5); -x_107 = lean_ctor_get(x_106, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_108 = x_106; -} else { - lean_dec_ref(x_106); - x_108 = lean_box(0); -} -if (lean_is_scalar(x_108)) { - x_109 = lean_alloc_ctor(0, 2, 0); -} else { - x_109 = x_108; -} -lean_ctor_set(x_109, 0, x_21); -lean_ctor_set(x_109, 1, x_107); -return x_109; +x_2 = x_5; +goto _start; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_95); -lean_ctor_set(x_110, 1, x_97); -x_111 = lean_st_ref_set(x_5, x_110, x_24); -lean_dec(x_5); -x_112 = lean_ctor_get(x_111, 1); -lean_inc(x_112); -if (lean_is_exclusive(x_111)) { - lean_ctor_release(x_111, 0); - lean_ctor_release(x_111, 1); - x_113 = x_111; -} else { - lean_dec_ref(x_111); - x_113 = lean_box(0); +uint8_t x_10; +x_10 = 1; +return x_10; } -if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(0, 2, 0); -} else { - x_114 = x_113; } -lean_ctor_set(x_114, 0, x_21); -lean_ctor_set(x_114, 1, x_112); -return x_114; } } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_115 = lean_box(0); -x_116 = lean_array_uset(x_76, x_92, x_115); -lean_inc(x_21); -x_117 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_21, x_93); -x_118 = lean_array_uset(x_116, x_92, x_117); -x_119 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_75); -lean_ctor_set(x_119, 1, x_118); -x_120 = lean_st_ref_set(x_5, x_119, x_24); -lean_dec(x_5); -x_121 = lean_ctor_get(x_120, 1); -lean_inc(x_121); -if (lean_is_exclusive(x_120)) { - lean_ctor_release(x_120, 0); - lean_ctor_release(x_120, 1); - x_122 = x_120; -} else { - lean_dec_ref(x_120); - x_122 = lean_box(0); -} -if (lean_is_scalar(x_122)) { - x_123 = lean_alloc_ctor(0, 2, 0); -} else { - x_123 = x_122; -} -lean_ctor_set(x_123, 0, x_21); -lean_ctor_set(x_123, 1, x_121); -return x_123; -} -} -} -} -else +if (lean_obj_tag(x_3) == 0) { -uint8_t x_129; -lean_dec(x_5); -lean_dec(x_2); lean_dec(x_1); -x_129 = !lean_is_exclusive(x_18); -if (x_129 == 0) -{ -return x_18; +return x_2; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_18, 0); -x_131 = lean_ctor_get(x_18, 1); -lean_inc(x_131); -lean_inc(x_130); -lean_dec(x_18); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -return x_132; -} -} -} -case 1: +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -lean_dec(x_4); -x_133 = lean_array_get_size(x_3); -x_134 = lean_unsigned_to_nat(0u); -x_135 = lean_unsigned_to_nat(1u); -x_136 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_133); -lean_ctor_set(x_136, 2, x_135); -x_137 = 0; -x_138 = lean_box(x_137); -x_139 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_139, 0, x_3); -lean_ctor_set(x_139, 1, x_138); -lean_inc(x_5); -x_140 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_136, x_136, x_139, x_134, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_136); -if (lean_obj_tag(x_140) == 0) -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_247; uint8_t x_248; -x_141 = lean_ctor_get(x_140, 0); -lean_inc(x_141); -x_142 = lean_ctor_get(x_140, 1); -lean_inc(x_142); -lean_dec(x_140); -x_247 = lean_ctor_get(x_141, 1); -lean_inc(x_247); -x_248 = lean_unbox(x_247); -lean_dec(x_247); -if (x_248 == 0) -{ -lean_dec(x_141); -lean_dec(x_2); +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); lean_inc(x_1); -x_143 = x_1; -goto block_246; -} -else -{ -lean_object* x_249; lean_object* x_250; -x_249 = lean_ctor_get(x_141, 0); -lean_inc(x_249); -lean_dec(x_141); -x_250 = l_Lean_mkAppN(x_2, x_249); -lean_dec(x_249); -x_143 = x_250; -goto block_246; -} -block_246: -{ -lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; -x_144 = lean_st_ref_take(x_5, x_142); -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_144, 1); -lean_inc(x_146); -lean_dec(x_144); -x_147 = !lean_is_exclusive(x_145); -if (x_147 == 0) -{ -lean_object* x_148; lean_object* x_149; lean_object* x_150; size_t x_151; uint64_t x_152; uint64_t x_153; uint64_t x_154; uint64_t x_155; uint64_t x_156; uint64_t x_157; uint64_t x_158; uint64_t x_159; uint64_t x_160; size_t x_161; size_t x_162; size_t x_163; size_t x_164; size_t x_165; lean_object* x_166; uint8_t x_167; -x_148 = lean_ctor_get(x_145, 0); -x_149 = lean_ctor_get(x_145, 1); -x_150 = lean_array_get_size(x_149); -x_151 = lean_ptr_addr(x_1); -x_152 = lean_usize_to_uint64(x_151); -x_153 = 11; -x_154 = lean_uint64_mix_hash(x_152, x_153); -x_155 = 32; -x_156 = lean_uint64_shift_right(x_154, x_155); -x_157 = lean_uint64_xor(x_154, x_156); -x_158 = 16; -x_159 = lean_uint64_shift_right(x_157, x_158); -x_160 = lean_uint64_xor(x_157, x_159); -x_161 = lean_uint64_to_usize(x_160); -x_162 = lean_usize_of_nat(x_150); -lean_dec(x_150); -x_163 = 1; -x_164 = lean_usize_sub(x_162, x_163); -x_165 = lean_usize_land(x_161, x_164); -x_166 = lean_array_uget(x_149, x_165); -x_167 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_166); -if (x_167 == 0) -{ -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; -x_168 = lean_nat_add(x_148, x_135); -lean_dec(x_148); -lean_inc(x_143); -x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_1); -lean_ctor_set(x_169, 1, x_143); -lean_ctor_set(x_169, 2, x_166); -x_170 = lean_array_uset(x_149, x_165, x_169); -x_171 = lean_unsigned_to_nat(4u); -x_172 = lean_nat_mul(x_168, x_171); -x_173 = lean_unsigned_to_nat(3u); -x_174 = lean_nat_div(x_172, x_173); -lean_dec(x_172); -x_175 = lean_array_get_size(x_170); -x_176 = lean_nat_dec_le(x_174, x_175); -lean_dec(x_175); -lean_dec(x_174); -if (x_176 == 0) -{ -lean_object* x_177; lean_object* x_178; uint8_t x_179; -x_177 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_170); -lean_ctor_set(x_145, 1, x_177); -lean_ctor_set(x_145, 0, x_168); -x_178 = lean_st_ref_set(x_5, x_145, x_146); -lean_dec(x_5); -x_179 = !lean_is_exclusive(x_178); -if (x_179 == 0) -{ -lean_object* x_180; -x_180 = lean_ctor_get(x_178, 0); -lean_dec(x_180); -lean_ctor_set(x_178, 0, x_143); -return x_178; -} -else -{ -lean_object* x_181; lean_object* x_182; -x_181 = lean_ctor_get(x_178, 1); -lean_inc(x_181); -lean_dec(x_178); -x_182 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_182, 0, x_143); -lean_ctor_set(x_182, 1, x_181); -return x_182; -} -} -else -{ -lean_object* x_183; uint8_t x_184; -lean_ctor_set(x_145, 1, x_170); -lean_ctor_set(x_145, 0, x_168); -x_183 = lean_st_ref_set(x_5, x_145, x_146); -lean_dec(x_5); -x_184 = !lean_is_exclusive(x_183); -if (x_184 == 0) -{ -lean_object* x_185; -x_185 = lean_ctor_get(x_183, 0); -lean_dec(x_185); -lean_ctor_set(x_183, 0, x_143); -return x_183; -} -else -{ -lean_object* x_186; lean_object* x_187; -x_186 = lean_ctor_get(x_183, 1); -lean_inc(x_186); -lean_dec(x_183); -x_187 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_187, 0, x_143); -lean_ctor_set(x_187, 1, x_186); -return x_187; -} -} -} -else -{ -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; uint8_t x_193; -x_188 = lean_box(0); -x_189 = lean_array_uset(x_149, x_165, x_188); -lean_inc(x_143); -x_190 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_143, x_166); -x_191 = lean_array_uset(x_189, x_165, x_190); -lean_ctor_set(x_145, 1, x_191); -x_192 = lean_st_ref_set(x_5, x_145, x_146); -lean_dec(x_5); -x_193 = !lean_is_exclusive(x_192); -if (x_193 == 0) -{ -lean_object* x_194; -x_194 = lean_ctor_get(x_192, 0); -lean_dec(x_194); -lean_ctor_set(x_192, 0, x_143); -return x_192; -} -else -{ -lean_object* x_195; lean_object* x_196; -x_195 = lean_ctor_get(x_192, 1); -lean_inc(x_195); -lean_dec(x_192); -x_196 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_196, 0, x_143); -lean_ctor_set(x_196, 1, x_195); -return x_196; -} -} -} -else -{ -lean_object* x_197; lean_object* x_198; lean_object* x_199; size_t x_200; uint64_t x_201; uint64_t x_202; uint64_t x_203; uint64_t x_204; uint64_t x_205; uint64_t x_206; uint64_t x_207; uint64_t x_208; uint64_t x_209; size_t x_210; size_t x_211; size_t x_212; size_t x_213; size_t x_214; lean_object* x_215; uint8_t x_216; -x_197 = lean_ctor_get(x_145, 0); -x_198 = lean_ctor_get(x_145, 1); -lean_inc(x_198); -lean_inc(x_197); -lean_dec(x_145); -x_199 = lean_array_get_size(x_198); -x_200 = lean_ptr_addr(x_1); -x_201 = lean_usize_to_uint64(x_200); -x_202 = 11; -x_203 = lean_uint64_mix_hash(x_201, x_202); -x_204 = 32; -x_205 = lean_uint64_shift_right(x_203, x_204); -x_206 = lean_uint64_xor(x_203, x_205); -x_207 = 16; -x_208 = lean_uint64_shift_right(x_206, x_207); -x_209 = lean_uint64_xor(x_206, x_208); -x_210 = lean_uint64_to_usize(x_209); -x_211 = lean_usize_of_nat(x_199); -lean_dec(x_199); -x_212 = 1; -x_213 = lean_usize_sub(x_211, x_212); -x_214 = lean_usize_land(x_210, x_213); -x_215 = lean_array_uget(x_198, x_214); -x_216 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_215); -if (x_216 == 0) -{ -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; uint8_t x_225; -x_217 = lean_nat_add(x_197, x_135); -lean_dec(x_197); -lean_inc(x_143); -x_218 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_218, 0, x_1); -lean_ctor_set(x_218, 1, x_143); -lean_ctor_set(x_218, 2, x_215); -x_219 = lean_array_uset(x_198, x_214, x_218); -x_220 = lean_unsigned_to_nat(4u); -x_221 = lean_nat_mul(x_217, x_220); -x_222 = lean_unsigned_to_nat(3u); -x_223 = lean_nat_div(x_221, x_222); -lean_dec(x_221); -x_224 = lean_array_get_size(x_219); -x_225 = lean_nat_dec_le(x_223, x_224); -lean_dec(x_224); -lean_dec(x_223); -if (x_225 == 0) -{ -lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_226 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_219); -x_227 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_227, 0, x_217); -lean_ctor_set(x_227, 1, x_226); -x_228 = lean_st_ref_set(x_5, x_227, x_146); -lean_dec(x_5); -x_229 = lean_ctor_get(x_228, 1); -lean_inc(x_229); -if (lean_is_exclusive(x_228)) { - lean_ctor_release(x_228, 0); - lean_ctor_release(x_228, 1); - x_230 = x_228; -} else { - lean_dec_ref(x_228); - x_230 = lean_box(0); -} -if (lean_is_scalar(x_230)) { - x_231 = lean_alloc_ctor(0, 2, 0); -} else { - x_231 = x_230; -} -lean_ctor_set(x_231, 0, x_143); -lean_ctor_set(x_231, 1, x_229); -return x_231; -} -else -{ -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_232 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_232, 0, x_217); -lean_ctor_set(x_232, 1, x_219); -x_233 = lean_st_ref_set(x_5, x_232, x_146); -lean_dec(x_5); -x_234 = lean_ctor_get(x_233, 1); -lean_inc(x_234); -if (lean_is_exclusive(x_233)) { - lean_ctor_release(x_233, 0); - lean_ctor_release(x_233, 1); - x_235 = x_233; -} else { - lean_dec_ref(x_233); - x_235 = lean_box(0); -} -if (lean_is_scalar(x_235)) { - x_236 = lean_alloc_ctor(0, 2, 0); -} else { - x_236 = x_235; -} -lean_ctor_set(x_236, 0, x_143); -lean_ctor_set(x_236, 1, x_234); -return x_236; -} -} -else -{ -lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; -x_237 = lean_box(0); -x_238 = lean_array_uset(x_198, x_214, x_237); -lean_inc(x_143); -x_239 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_143, x_215); -x_240 = lean_array_uset(x_238, x_214, x_239); -x_241 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_241, 0, x_197); -lean_ctor_set(x_241, 1, x_240); -x_242 = lean_st_ref_set(x_5, x_241, x_146); -lean_dec(x_5); -x_243 = lean_ctor_get(x_242, 1); -lean_inc(x_243); -if (lean_is_exclusive(x_242)) { - lean_ctor_release(x_242, 0); - lean_ctor_release(x_242, 1); - x_244 = x_242; -} else { - lean_dec_ref(x_242); - x_244 = lean_box(0); -} -if (lean_is_scalar(x_244)) { - x_245 = lean_alloc_ctor(0, 2, 0); -} else { - x_245 = x_244; -} -lean_ctor_set(x_245, 0, x_143); -lean_ctor_set(x_245, 1, x_243); -return x_245; -} -} -} -} -else -{ -uint8_t x_251; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_251 = !lean_is_exclusive(x_140); -if (x_251 == 0) -{ -return x_140; -} -else -{ -lean_object* x_252; lean_object* x_253; lean_object* x_254; -x_252 = lean_ctor_get(x_140, 0); -x_253 = lean_ctor_get(x_140, 1); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_140); -x_254 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_253); -return x_254; -} -} -} -case 2: -{ -lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -lean_dec(x_4); -x_255 = lean_array_get_size(x_3); -x_256 = lean_unsigned_to_nat(0u); -x_257 = lean_unsigned_to_nat(1u); -x_258 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_258, 0, x_256); -lean_ctor_set(x_258, 1, x_255); -lean_ctor_set(x_258, 2, x_257); -x_259 = 0; -x_260 = lean_box(x_259); -x_261 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_261, 0, x_3); -lean_ctor_set(x_261, 1, x_260); lean_inc(x_5); -x_262 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_258, x_258, x_261, x_256, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_258); -if (lean_obj_tag(x_262) == 0) -{ -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_369; uint8_t x_370; -x_263 = lean_ctor_get(x_262, 0); -lean_inc(x_263); -x_264 = lean_ctor_get(x_262, 1); -lean_inc(x_264); -lean_dec(x_262); -x_369 = lean_ctor_get(x_263, 1); -lean_inc(x_369); -x_370 = lean_unbox(x_369); -lean_dec(x_369); -if (x_370 == 0) -{ -lean_dec(x_263); -lean_dec(x_2); -lean_inc(x_1); -x_265 = x_1; -goto block_368; -} -else -{ -lean_object* x_371; lean_object* x_372; -x_371 = lean_ctor_get(x_263, 0); -lean_inc(x_371); -lean_dec(x_263); -x_372 = l_Lean_mkAppN(x_2, x_371); -lean_dec(x_371); -x_265 = x_372; -goto block_368; -} -block_368: -{ -lean_object* x_266; lean_object* x_267; lean_object* x_268; uint8_t x_269; -x_266 = lean_st_ref_take(x_5, x_264); -x_267 = lean_ctor_get(x_266, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_266, 1); -lean_inc(x_268); -lean_dec(x_266); -x_269 = !lean_is_exclusive(x_267); -if (x_269 == 0) -{ -lean_object* x_270; lean_object* x_271; lean_object* x_272; size_t x_273; uint64_t x_274; uint64_t x_275; uint64_t x_276; uint64_t x_277; uint64_t x_278; uint64_t x_279; uint64_t x_280; uint64_t x_281; uint64_t x_282; size_t x_283; size_t x_284; size_t x_285; size_t x_286; size_t x_287; lean_object* x_288; uint8_t x_289; -x_270 = lean_ctor_get(x_267, 0); -x_271 = lean_ctor_get(x_267, 1); -x_272 = lean_array_get_size(x_271); -x_273 = lean_ptr_addr(x_1); -x_274 = lean_usize_to_uint64(x_273); -x_275 = 11; -x_276 = lean_uint64_mix_hash(x_274, x_275); -x_277 = 32; -x_278 = lean_uint64_shift_right(x_276, x_277); -x_279 = lean_uint64_xor(x_276, x_278); -x_280 = 16; -x_281 = lean_uint64_shift_right(x_279, x_280); -x_282 = lean_uint64_xor(x_279, x_281); -x_283 = lean_uint64_to_usize(x_282); -x_284 = lean_usize_of_nat(x_272); -lean_dec(x_272); -x_285 = 1; -x_286 = lean_usize_sub(x_284, x_285); -x_287 = lean_usize_land(x_283, x_286); -x_288 = lean_array_uget(x_271, x_287); -x_289 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_288); -if (x_289 == 0) -{ -lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; uint8_t x_298; -x_290 = lean_nat_add(x_270, x_257); -lean_dec(x_270); -lean_inc(x_265); -x_291 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_291, 0, x_1); -lean_ctor_set(x_291, 1, x_265); -lean_ctor_set(x_291, 2, x_288); -x_292 = lean_array_uset(x_271, x_287, x_291); -x_293 = lean_unsigned_to_nat(4u); -x_294 = lean_nat_mul(x_290, x_293); -x_295 = lean_unsigned_to_nat(3u); -x_296 = lean_nat_div(x_294, x_295); -lean_dec(x_294); -x_297 = lean_array_get_size(x_292); -x_298 = lean_nat_dec_le(x_296, x_297); -lean_dec(x_297); -lean_dec(x_296); -if (x_298 == 0) -{ -lean_object* x_299; lean_object* x_300; uint8_t x_301; -x_299 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_292); -lean_ctor_set(x_267, 1, x_299); -lean_ctor_set(x_267, 0, x_290); -x_300 = lean_st_ref_set(x_5, x_267, x_268); -lean_dec(x_5); -x_301 = !lean_is_exclusive(x_300); -if (x_301 == 0) -{ -lean_object* x_302; -x_302 = lean_ctor_get(x_300, 0); -lean_dec(x_302); -lean_ctor_set(x_300, 0, x_265); -return x_300; -} -else -{ -lean_object* x_303; lean_object* x_304; -x_303 = lean_ctor_get(x_300, 1); -lean_inc(x_303); -lean_dec(x_300); -x_304 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_304, 0, x_265); -lean_ctor_set(x_304, 1, x_303); -return x_304; -} -} -else -{ -lean_object* x_305; uint8_t x_306; -lean_ctor_set(x_267, 1, x_292); -lean_ctor_set(x_267, 0, x_290); -x_305 = lean_st_ref_set(x_5, x_267, x_268); -lean_dec(x_5); -x_306 = !lean_is_exclusive(x_305); -if (x_306 == 0) -{ -lean_object* x_307; -x_307 = lean_ctor_get(x_305, 0); -lean_dec(x_307); -lean_ctor_set(x_305, 0, x_265); -return x_305; +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = 32; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = 16; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_18 = 1; +x_19 = lean_usize_sub(x_17, x_18); +x_20 = lean_usize_land(x_16, x_19); +x_21 = lean_array_uget(x_2, x_20); +lean_ctor_set(x_3, 2, x_21); +x_22 = lean_array_uset(x_2, x_20, x_3); +x_2 = x_22; +x_3 = x_6; +goto _start; } else { -lean_object* x_308; lean_object* x_309; -x_308 = lean_ctor_get(x_305, 1); -lean_inc(x_308); -lean_dec(x_305); -x_309 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_309, 0, x_265); -lean_ctor_set(x_309, 1, x_308); -return x_309; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_27 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_24); +x_28 = lean_apply_1(x_1, x_24); +x_29 = lean_unbox_uint64(x_28); +lean_dec(x_28); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_2, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_uset(x_2, x_40, x_42); +x_2 = x_43; +x_3 = x_26; +goto _start; } } } -else +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; uint8_t x_315; -x_310 = lean_box(0); -x_311 = lean_array_uset(x_271, x_287, x_310); -lean_inc(x_265); -x_312 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_265, x_288); -x_313 = lean_array_uset(x_311, x_287, x_312); -lean_ctor_set(x_267, 1, x_313); -x_314 = lean_st_ref_set(x_5, x_267, x_268); -lean_dec(x_5); -x_315 = !lean_is_exclusive(x_314); -if (x_315 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_316; -x_316 = lean_ctor_get(x_314, 0); -lean_dec(x_316); -lean_ctor_set(x_314, 0, x_265); -return x_314; +return x_1; } else { -lean_object* x_317; lean_object* x_318; -x_317 = lean_ctor_get(x_314, 1); -lean_inc(x_317); -lean_dec(x_314); -x_318 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_318, 0, x_265); -lean_ctor_set(x_318, 1, x_317); -return x_318; -} -} -} -else +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) { -lean_object* x_319; lean_object* x_320; lean_object* x_321; size_t x_322; uint64_t x_323; uint64_t x_324; uint64_t x_325; uint64_t x_326; uint64_t x_327; uint64_t x_328; uint64_t x_329; uint64_t x_330; uint64_t x_331; size_t x_332; size_t x_333; size_t x_334; size_t x_335; size_t x_336; lean_object* x_337; uint8_t x_338; -x_319 = lean_ctor_get(x_267, 0); -x_320 = lean_ctor_get(x_267, 1); -lean_inc(x_320); -lean_inc(x_319); -lean_dec(x_267); -x_321 = lean_array_get_size(x_320); -x_322 = lean_ptr_addr(x_1); -x_323 = lean_usize_to_uint64(x_322); -x_324 = 11; -x_325 = lean_uint64_mix_hash(x_323, x_324); -x_326 = 32; -x_327 = lean_uint64_shift_right(x_325, x_326); -x_328 = lean_uint64_xor(x_325, x_327); -x_329 = 16; -x_330 = lean_uint64_shift_right(x_328, x_329); -x_331 = lean_uint64_xor(x_328, x_330); -x_332 = lean_uint64_to_usize(x_331); -x_333 = lean_usize_of_nat(x_321); -lean_dec(x_321); -x_334 = 1; -x_335 = lean_usize_sub(x_333, x_334); -x_336 = lean_usize_land(x_332, x_335); -x_337 = lean_array_uget(x_320, x_336); -x_338 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_337); -if (x_338 == 0) -{ -lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; uint8_t x_347; -x_339 = lean_nat_add(x_319, x_257); -lean_dec(x_319); -lean_inc(x_265); -x_340 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_340, 0, x_1); -lean_ctor_set(x_340, 1, x_265); -lean_ctor_set(x_340, 2, x_337); -x_341 = lean_array_uset(x_320, x_336, x_340); -x_342 = lean_unsigned_to_nat(4u); -x_343 = lean_nat_mul(x_339, x_342); -x_344 = lean_unsigned_to_nat(3u); -x_345 = lean_nat_div(x_343, x_344); -lean_dec(x_343); -x_346 = lean_array_get_size(x_341); -x_347 = lean_nat_dec_le(x_345, x_346); -lean_dec(x_346); -lean_dec(x_345); -if (x_347 == 0) -{ -lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -x_348 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_341); -x_349 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_349, 0, x_339); -lean_ctor_set(x_349, 1, x_348); -x_350 = lean_st_ref_set(x_5, x_349, x_268); -lean_dec(x_5); -x_351 = lean_ctor_get(x_350, 1); -lean_inc(x_351); -if (lean_is_exclusive(x_350)) { - lean_ctor_release(x_350, 0); - lean_ctor_release(x_350, 1); - x_352 = x_350; -} else { - lean_dec_ref(x_350); - x_352 = lean_box(0); -} -if (lean_is_scalar(x_352)) { - x_353 = lean_alloc_ctor(0, 2, 0); -} else { - x_353 = x_352; -} -lean_ctor_set(x_353, 0, x_265); -lean_ctor_set(x_353, 1, x_351); -return x_353; +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_usize_to_uint64(x_7); +x_9 = 11; +x_10 = lean_uint64_mix_hash(x_8, x_9); +x_11 = 32; +x_12 = lean_uint64_shift_right(x_10, x_11); +x_13 = lean_uint64_xor(x_10, x_12); +x_14 = 16; +x_15 = lean_uint64_shift_right(x_13, x_14); +x_16 = lean_uint64_xor(x_13, x_15); +x_17 = lean_uint64_to_usize(x_16); +x_18 = lean_usize_of_nat(x_6); +lean_dec(x_6); +x_19 = 1; +x_20 = lean_usize_sub(x_18, x_19); +x_21 = lean_usize_land(x_17, x_20); +x_22 = lean_array_uget(x_1, x_21); +lean_ctor_set(x_2, 2, x_22); +x_23 = lean_array_uset(x_1, x_21, x_2); +x_1 = x_23; +x_2 = x_5; +goto _start; } else { -lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; -x_354 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_354, 0, x_339); -lean_ctor_set(x_354, 1, x_341); -x_355 = lean_st_ref_set(x_5, x_354, x_268); -lean_dec(x_5); -x_356 = lean_ctor_get(x_355, 1); -lean_inc(x_356); -if (lean_is_exclusive(x_355)) { - lean_ctor_release(x_355, 0); - lean_ctor_release(x_355, 1); - x_357 = x_355; -} else { - lean_dec_ref(x_355); - x_357 = lean_box(0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; size_t x_41; size_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_25 = lean_ctor_get(x_2, 0); +x_26 = lean_ctor_get(x_2, 1); +x_27 = lean_ctor_get(x_2, 2); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_2); +x_28 = lean_array_get_size(x_1); +x_29 = lean_ptr_addr(x_25); +x_30 = lean_usize_to_uint64(x_29); +x_31 = 11; +x_32 = lean_uint64_mix_hash(x_30, x_31); +x_33 = 32; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = 16; +x_37 = lean_uint64_shift_right(x_35, x_36); +x_38 = lean_uint64_xor(x_35, x_37); +x_39 = lean_uint64_to_usize(x_38); +x_40 = lean_usize_of_nat(x_28); +lean_dec(x_28); +x_41 = 1; +x_42 = lean_usize_sub(x_40, x_41); +x_43 = lean_usize_land(x_39, x_42); +x_44 = lean_array_uget(x_1, x_43); +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_25); +lean_ctor_set(x_45, 1, x_26); +lean_ctor_set(x_45, 2, x_44); +x_46 = lean_array_uset(x_1, x_43, x_45); +x_1 = x_46; +x_2 = x_27; +goto _start; } -if (lean_is_scalar(x_357)) { - x_358 = lean_alloc_ctor(0, 2, 0); -} else { - x_358 = x_357; } -lean_ctor_set(x_358, 0, x_265); -lean_ctor_set(x_358, 1, x_356); -return x_358; } } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; -x_359 = lean_box(0); -x_360 = lean_array_uset(x_320, x_336, x_359); -lean_inc(x_265); -x_361 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_265, x_337); -x_362 = lean_array_uset(x_360, x_336, x_361); -x_363 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_363, 0, x_319); -lean_ctor_set(x_363, 1, x_362); -x_364 = lean_st_ref_set(x_5, x_363, x_268); -lean_dec(x_5); -x_365 = lean_ctor_get(x_364, 1); -lean_inc(x_365); -if (lean_is_exclusive(x_364)) { - lean_ctor_release(x_364, 0); - lean_ctor_release(x_364, 1); - x_366 = x_364; -} else { - lean_dec_ref(x_364); - x_366 = lean_box(0); -} -if (lean_is_scalar(x_366)) { - x_367 = lean_alloc_ctor(0, 2, 0); -} else { - x_367 = x_366; -} -lean_ctor_set(x_367, 0, x_265); -lean_ctor_set(x_367, 1, x_365); -return x_367; -} -} -} -} -else +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) { -uint8_t x_373; -lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_373 = !lean_is_exclusive(x_262); -if (x_373 == 0) -{ -return x_262; +return x_3; } else { -lean_object* x_374; lean_object* x_375; lean_object* x_376; -x_374 = lean_ctor_get(x_262, 0); -x_375 = lean_ctor_get(x_262, 1); -lean_inc(x_375); -lean_inc(x_374); -lean_dec(x_262); -x_376 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_376, 0, x_374); -lean_ctor_set(x_376, 1, x_375); -return x_376; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__5(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; } } } -case 3: +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(lean_object* x_1) { +_start: { -lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; uint8_t x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; -lean_dec(x_4); -x_377 = lean_array_get_size(x_3); -x_378 = lean_unsigned_to_nat(0u); -x_379 = lean_unsigned_to_nat(1u); -x_380 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_380, 0, x_378); -lean_ctor_set(x_380, 1, x_377); -lean_ctor_set(x_380, 2, x_379); -x_381 = 0; -x_382 = lean_box(x_381); -x_383 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_383, 0, x_3); -lean_ctor_set(x_383, 1, x_382); -lean_inc(x_5); -x_384 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_380, x_380, x_383, x_378, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_380); -if (lean_obj_tag(x_384) == 0) -{ -lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_491; uint8_t x_492; -x_385 = lean_ctor_get(x_384, 0); -lean_inc(x_385); -x_386 = lean_ctor_get(x_384, 1); -lean_inc(x_386); -lean_dec(x_384); -x_491 = lean_ctor_get(x_385, 1); -lean_inc(x_491); -x_492 = lean_unbox(x_491); -lean_dec(x_491); -if (x_492 == 0) -{ -lean_dec(x_385); +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); lean_dec(x_2); -lean_inc(x_1); -x_387 = x_1; -goto block_490; -} -else -{ -lean_object* x_493; lean_object* x_494; -x_493 = lean_ctor_get(x_385, 0); -lean_inc(x_493); -lean_dec(x_385); -x_494 = l_Lean_mkAppN(x_2, x_493); -lean_dec(x_493); -x_387 = x_494; -goto block_490; -} -block_490: -{ -lean_object* x_388; lean_object* x_389; lean_object* x_390; uint8_t x_391; -x_388 = lean_st_ref_take(x_5, x_386); -x_389 = lean_ctor_get(x_388, 0); -lean_inc(x_389); -x_390 = lean_ctor_get(x_388, 1); -lean_inc(x_390); -lean_dec(x_388); -x_391 = !lean_is_exclusive(x_389); -if (x_391 == 0) -{ -lean_object* x_392; lean_object* x_393; lean_object* x_394; size_t x_395; uint64_t x_396; uint64_t x_397; uint64_t x_398; uint64_t x_399; uint64_t x_400; uint64_t x_401; uint64_t x_402; uint64_t x_403; uint64_t x_404; size_t x_405; size_t x_406; size_t x_407; size_t x_408; size_t x_409; lean_object* x_410; uint8_t x_411; -x_392 = lean_ctor_get(x_389, 0); -x_393 = lean_ctor_get(x_389, 1); -x_394 = lean_array_get_size(x_393); -x_395 = lean_ptr_addr(x_1); -x_396 = lean_usize_to_uint64(x_395); -x_397 = 11; -x_398 = lean_uint64_mix_hash(x_396, x_397); -x_399 = 32; -x_400 = lean_uint64_shift_right(x_398, x_399); -x_401 = lean_uint64_xor(x_398, x_400); -x_402 = 16; -x_403 = lean_uint64_shift_right(x_401, x_402); -x_404 = lean_uint64_xor(x_401, x_403); -x_405 = lean_uint64_to_usize(x_404); -x_406 = lean_usize_of_nat(x_394); -lean_dec(x_394); -x_407 = 1; -x_408 = lean_usize_sub(x_406, x_407); -x_409 = lean_usize_land(x_405, x_408); -x_410 = lean_array_uget(x_393, x_409); -x_411 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_410); -if (x_411 == 0) -{ -lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; uint8_t x_420; -x_412 = lean_nat_add(x_392, x_379); -lean_dec(x_392); -lean_inc(x_387); -x_413 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_413, 0, x_1); -lean_ctor_set(x_413, 1, x_387); -lean_ctor_set(x_413, 2, x_410); -x_414 = lean_array_uset(x_393, x_409, x_413); -x_415 = lean_unsigned_to_nat(4u); -x_416 = lean_nat_mul(x_412, x_415); -x_417 = lean_unsigned_to_nat(3u); -x_418 = lean_nat_div(x_416, x_417); -lean_dec(x_416); -x_419 = lean_array_get_size(x_414); -x_420 = lean_nat_dec_le(x_418, x_419); -lean_dec(x_419); -lean_dec(x_418); -if (x_420 == 0) -{ -lean_object* x_421; lean_object* x_422; uint8_t x_423; -x_421 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_414); -lean_ctor_set(x_389, 1, x_421); -lean_ctor_set(x_389, 0, x_412); -x_422 = lean_st_ref_set(x_5, x_389, x_390); -lean_dec(x_5); -x_423 = !lean_is_exclusive(x_422); -if (x_423 == 0) -{ -lean_object* x_424; -x_424 = lean_ctor_get(x_422, 0); -lean_dec(x_424); -lean_ctor_set(x_422, 0, x_387); -return x_422; -} -else -{ -lean_object* x_425; lean_object* x_426; -x_425 = lean_ctor_get(x_422, 1); -lean_inc(x_425); -lean_dec(x_422); -x_426 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_426, 0, x_387); -lean_ctor_set(x_426, 1, x_425); -return x_426; +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_7, x_1, x_6); +return x_8; } } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_427; uint8_t x_428; -lean_ctor_set(x_389, 1, x_414); -lean_ctor_set(x_389, 0, x_412); -x_427 = lean_st_ref_set(x_5, x_389, x_390); -lean_dec(x_5); -x_428 = !lean_is_exclusive(x_427); -if (x_428 == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_429; -x_429 = lean_ctor_get(x_427, 0); -lean_dec(x_429); -lean_ctor_set(x_427, 0, x_387); -return x_427; +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; } else { -lean_object* x_430; lean_object* x_431; -x_430 = lean_ctor_get(x_427, 1); -lean_inc(x_430); -lean_dec(x_427); -x_431 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_431, 0, x_387); -lean_ctor_set(x_431, 1, x_430); -return x_431; -} -} -} -else +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; uint8_t x_437; -x_432 = lean_box(0); -x_433 = lean_array_uset(x_393, x_409, x_432); -lean_inc(x_387); -x_434 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_387, x_410); -x_435 = lean_array_uset(x_433, x_409, x_434); -lean_ctor_set(x_389, 1, x_435); -x_436 = lean_st_ref_set(x_5, x_389, x_390); -lean_dec(x_5); -x_437 = !lean_is_exclusive(x_436); -if (x_437 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; uint8_t x_11; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_ptr_addr(x_6); +x_10 = lean_ptr_addr(x_1); +x_11 = lean_usize_dec_eq(x_9, x_10); +if (x_11 == 0) { -lean_object* x_438; -x_438 = lean_ctor_get(x_436, 0); -lean_dec(x_438); -lean_ctor_set(x_436, 0, x_387); -return x_436; +lean_object* x_12; +x_12 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_12); +return x_3; } else { -lean_object* x_439; lean_object* x_440; -x_439 = lean_ctor_get(x_436, 1); -lean_inc(x_439); -lean_dec(x_436); -x_440 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_440, 0, x_387); -lean_ctor_set(x_440, 1, x_439); -return x_440; -} +lean_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } else { -lean_object* x_441; lean_object* x_442; lean_object* x_443; size_t x_444; uint64_t x_445; uint64_t x_446; uint64_t x_447; uint64_t x_448; uint64_t x_449; uint64_t x_450; uint64_t x_451; uint64_t x_452; uint64_t x_453; size_t x_454; size_t x_455; size_t x_456; size_t x_457; size_t x_458; lean_object* x_459; uint8_t x_460; -x_441 = lean_ctor_get(x_389, 0); -x_442 = lean_ctor_get(x_389, 1); -lean_inc(x_442); -lean_inc(x_441); -lean_dec(x_389); -x_443 = lean_array_get_size(x_442); -x_444 = lean_ptr_addr(x_1); -x_445 = lean_usize_to_uint64(x_444); -x_446 = 11; -x_447 = lean_uint64_mix_hash(x_445, x_446); -x_448 = 32; -x_449 = lean_uint64_shift_right(x_447, x_448); -x_450 = lean_uint64_xor(x_447, x_449); -x_451 = 16; -x_452 = lean_uint64_shift_right(x_450, x_451); -x_453 = lean_uint64_xor(x_450, x_452); -x_454 = lean_uint64_to_usize(x_453); -x_455 = lean_usize_of_nat(x_443); -lean_dec(x_443); -x_456 = 1; -x_457 = lean_usize_sub(x_455, x_456); -x_458 = lean_usize_land(x_454, x_457); -x_459 = lean_array_uget(x_442, x_458); -x_460 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_459); -if (x_460 == 0) -{ -lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; uint8_t x_469; -x_461 = lean_nat_add(x_441, x_379); -lean_dec(x_441); -lean_inc(x_387); -x_462 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_462, 0, x_1); -lean_ctor_set(x_462, 1, x_387); -lean_ctor_set(x_462, 2, x_459); -x_463 = lean_array_uset(x_442, x_458, x_462); -x_464 = lean_unsigned_to_nat(4u); -x_465 = lean_nat_mul(x_461, x_464); -x_466 = lean_unsigned_to_nat(3u); -x_467 = lean_nat_div(x_465, x_466); -lean_dec(x_465); -x_468 = lean_array_get_size(x_463); -x_469 = lean_nat_dec_le(x_467, x_468); -lean_dec(x_468); -lean_dec(x_467); -if (x_469 == 0) -{ -lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; -x_470 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_463); -x_471 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_471, 0, x_461); -lean_ctor_set(x_471, 1, x_470); -x_472 = lean_st_ref_set(x_5, x_471, x_390); -lean_dec(x_5); -x_473 = lean_ctor_get(x_472, 1); -lean_inc(x_473); -if (lean_is_exclusive(x_472)) { - lean_ctor_release(x_472, 0); - lean_ctor_release(x_472, 1); - x_474 = x_472; -} else { - lean_dec_ref(x_472); - x_474 = lean_box(0); -} -if (lean_is_scalar(x_474)) { - x_475 = lean_alloc_ctor(0, 2, 0); -} else { - x_475 = x_474; -} -lean_ctor_set(x_475, 0, x_387); -lean_ctor_set(x_475, 1, x_473); -return x_475; -} -else +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; uint8_t x_18; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_3, 2); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_16 = lean_ptr_addr(x_13); +x_17 = lean_ptr_addr(x_1); +x_18 = lean_usize_dec_eq(x_16, x_17); +if (x_18 == 0) { -lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; -x_476 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_476, 0, x_461); -lean_ctor_set(x_476, 1, x_463); -x_477 = lean_st_ref_set(x_5, x_476, x_390); -lean_dec(x_5); -x_478 = lean_ctor_get(x_477, 1); -lean_inc(x_478); -if (lean_is_exclusive(x_477)) { - lean_ctor_release(x_477, 0); - lean_ctor_release(x_477, 1); - x_479 = x_477; -} else { - lean_dec_ref(x_477); - x_479 = lean_box(0); -} -if (lean_is_scalar(x_479)) { - x_480 = lean_alloc_ctor(0, 2, 0); -} else { - x_480 = x_479; -} -lean_ctor_set(x_480, 0, x_387); -lean_ctor_set(x_480, 1, x_478); -return x_480; -} +lean_object* x_19; lean_object* x_20; +x_19 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_2, x_15); +x_20 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +lean_ctor_set(x_20, 2, x_19); +return x_20; } else { -lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; -x_481 = lean_box(0); -x_482 = lean_array_uset(x_442, x_458, x_481); -lean_inc(x_387); -x_483 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_387, x_459); -x_484 = lean_array_uset(x_482, x_458, x_483); -x_485 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_485, 0, x_441); -lean_ctor_set(x_485, 1, x_484); -x_486 = lean_st_ref_set(x_5, x_485, x_390); -lean_dec(x_5); -x_487 = lean_ctor_get(x_486, 1); -lean_inc(x_487); -if (lean_is_exclusive(x_486)) { - lean_ctor_release(x_486, 0); - lean_ctor_release(x_486, 1); - x_488 = x_486; -} else { - lean_dec_ref(x_486); - x_488 = lean_box(0); -} -if (lean_is_scalar(x_488)) { - x_489 = lean_alloc_ctor(0, 2, 0); -} else { - x_489 = x_488; +lean_object* x_21; +lean_dec(x_14); +lean_dec(x_13); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_15); +return x_21; } -lean_ctor_set(x_489, 0, x_387); -lean_ctor_set(x_489, 1, x_487); -return x_489; } } } } -else -{ -uint8_t x_495; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_495 = !lean_is_exclusive(x_384); -if (x_495 == 0) +static lean_object* _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1() { +_start: { -return x_384; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} } -else +static lean_object* _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2() { +_start: { -lean_object* x_496; lean_object* x_497; lean_object* x_498; -x_496 = lean_ctor_get(x_384, 0); -x_497 = lean_ctor_get(x_384, 1); -lean_inc(x_497); -lean_inc(x_496); -lean_dec(x_384); -x_498 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_498, 0, x_496); -lean_ctor_set(x_498, 1, x_497); -return x_498; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1; +x_2 = l_Lean_instInhabitedExpr; +x_3 = l_instInhabitedOfMonad___rarg(x_1, x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2; +x_9 = lean_panic_fn(x_8, x_1); +x_10 = lean_apply_6(x_9, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; +} } -case 4: +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_2, 1); +x_14 = lean_nat_dec_lt(x_4, x_13); +if (x_14 == 0) { -lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; uint8_t x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; +lean_object* x_15; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_4); -x_499 = lean_array_get_size(x_3); -x_500 = lean_unsigned_to_nat(0u); -x_501 = lean_unsigned_to_nat(1u); -x_502 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_502, 0, x_500); -lean_ctor_set(x_502, 1, x_499); -lean_ctor_set(x_502, 2, x_501); -x_503 = 0; -x_504 = lean_box(x_503); -x_505 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_505, 0, x_3); -lean_ctor_set(x_505, 1, x_504); -lean_inc(x_5); -x_506 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_502, x_502, x_505, x_500, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_502); -if (lean_obj_tag(x_506) == 0) -{ -lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_613; uint8_t x_614; -x_507 = lean_ctor_get(x_506, 0); -lean_inc(x_507); -x_508 = lean_ctor_get(x_506, 1); -lean_inc(x_508); -lean_dec(x_506); -x_613 = lean_ctor_get(x_507, 1); -lean_inc(x_613); -x_614 = lean_unbox(x_613); -lean_dec(x_613); -if (x_614 == 0) -{ -lean_dec(x_507); -lean_dec(x_2); -lean_inc(x_1); -x_509 = x_1; -goto block_612; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_3); +lean_ctor_set(x_15, 1, x_12); +return x_15; } else { -lean_object* x_615; lean_object* x_616; -x_615 = lean_ctor_get(x_507, 0); -lean_inc(x_615); -lean_dec(x_507); -x_616 = l_Lean_mkAppN(x_2, x_615); -lean_dec(x_615); -x_509 = x_616; -goto block_612; -} -block_612: -{ -lean_object* x_510; lean_object* x_511; lean_object* x_512; uint8_t x_513; -x_510 = lean_st_ref_take(x_5, x_508); -x_511 = lean_ctor_get(x_510, 0); -lean_inc(x_511); -x_512 = lean_ctor_get(x_510, 1); -lean_inc(x_512); -lean_dec(x_510); -x_513 = !lean_is_exclusive(x_511); -if (x_513 == 0) -{ -lean_object* x_514; lean_object* x_515; lean_object* x_516; size_t x_517; uint64_t x_518; uint64_t x_519; uint64_t x_520; uint64_t x_521; uint64_t x_522; uint64_t x_523; uint64_t x_524; uint64_t x_525; uint64_t x_526; size_t x_527; size_t x_528; size_t x_529; size_t x_530; size_t x_531; lean_object* x_532; uint8_t x_533; -x_514 = lean_ctor_get(x_511, 0); -x_515 = lean_ctor_get(x_511, 1); -x_516 = lean_array_get_size(x_515); -x_517 = lean_ptr_addr(x_1); -x_518 = lean_usize_to_uint64(x_517); -x_519 = 11; -x_520 = lean_uint64_mix_hash(x_518, x_519); -x_521 = 32; -x_522 = lean_uint64_shift_right(x_520, x_521); -x_523 = lean_uint64_xor(x_520, x_522); -x_524 = 16; -x_525 = lean_uint64_shift_right(x_523, x_524); -x_526 = lean_uint64_xor(x_523, x_525); -x_527 = lean_uint64_to_usize(x_526); -x_528 = lean_usize_of_nat(x_516); -lean_dec(x_516); -x_529 = 1; -x_530 = lean_usize_sub(x_528, x_529); -x_531 = lean_usize_land(x_527, x_530); -x_532 = lean_array_uget(x_515, x_531); -x_533 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_532); -if (x_533 == 0) -{ -lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; uint8_t x_542; -x_534 = lean_nat_add(x_514, x_501); -lean_dec(x_514); -lean_inc(x_509); -x_535 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_535, 0, x_1); -lean_ctor_set(x_535, 1, x_509); -lean_ctor_set(x_535, 2, x_532); -x_536 = lean_array_uset(x_515, x_531, x_535); -x_537 = lean_unsigned_to_nat(4u); -x_538 = lean_nat_mul(x_534, x_537); -x_539 = lean_unsigned_to_nat(3u); -x_540 = lean_nat_div(x_538, x_539); -lean_dec(x_538); -x_541 = lean_array_get_size(x_536); -x_542 = lean_nat_dec_le(x_540, x_541); -lean_dec(x_541); -lean_dec(x_540); -if (x_542 == 0) -{ -lean_object* x_543; lean_object* x_544; uint8_t x_545; -x_543 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_536); -lean_ctor_set(x_511, 1, x_543); -lean_ctor_set(x_511, 0, x_534); -x_544 = lean_st_ref_set(x_5, x_511, x_512); -lean_dec(x_5); -x_545 = !lean_is_exclusive(x_544); -if (x_545 == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_23; +x_23 = !lean_is_exclusive(x_3); +if (x_23 == 0) { -lean_object* x_546; -x_546 = lean_ctor_get(x_544, 0); -lean_dec(x_546); -lean_ctor_set(x_544, 0, x_509); -return x_544; -} -else +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_array_get_size(x_24); +x_27 = lean_nat_dec_lt(x_4, x_26); +lean_dec(x_26); +if (x_27 == 0) { -lean_object* x_547; lean_object* x_548; -x_547 = lean_ctor_get(x_544, 1); -lean_inc(x_547); -lean_dec(x_544); -x_548 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_548, 0, x_509); -lean_ctor_set(x_548, 1, x_547); -return x_548; -} -} -else +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = l_Lean_instInhabitedExpr; +x_29 = l_outOfBounds___rarg(x_28); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_29); +x_30 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_29, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_549; uint8_t x_550; -lean_ctor_set(x_511, 1, x_536); -lean_ctor_set(x_511, 0, x_534); -x_549 = lean_st_ref_set(x_5, x_511, x_512); -lean_dec(x_5); -x_550 = !lean_is_exclusive(x_549); -if (x_550 == 0) +lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; uint8_t x_35; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_ptr_addr(x_29); +lean_dec(x_29); +x_34 = lean_ptr_addr(x_31); +x_35 = lean_usize_dec_eq(x_33, x_34); +if (x_35 == 0) { -lean_object* x_551; -x_551 = lean_ctor_get(x_549, 0); -lean_dec(x_551); -lean_ctor_set(x_549, 0, x_509); -return x_549; +lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_25); +x_36 = lean_array_set(x_24, x_4, x_31); +x_37 = 1; +x_38 = lean_box(x_37); +lean_ctor_set(x_3, 1, x_38); +lean_ctor_set(x_3, 0, x_36); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_3); +x_16 = x_39; +x_17 = x_32; +goto block_22; } else { -lean_object* x_552; lean_object* x_553; -x_552 = lean_ctor_get(x_549, 1); -lean_inc(x_552); -lean_dec(x_549); -x_553 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_553, 0, x_509); -lean_ctor_set(x_553, 1, x_552); -return x_553; -} +lean_object* x_40; +lean_dec(x_31); +x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_40, 0, x_3); +x_16 = x_40; +x_17 = x_32; +goto block_22; } } else { -lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; uint8_t x_559; -x_554 = lean_box(0); -x_555 = lean_array_uset(x_515, x_531, x_554); -lean_inc(x_509); -x_556 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_509, x_532); -x_557 = lean_array_uset(x_555, x_531, x_556); -lean_ctor_set(x_511, 1, x_557); -x_558 = lean_st_ref_set(x_5, x_511, x_512); -lean_dec(x_5); -x_559 = !lean_is_exclusive(x_558); -if (x_559 == 0) +uint8_t x_41; +lean_dec(x_29); +lean_free_object(x_3); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_41 = !lean_is_exclusive(x_30); +if (x_41 == 0) { -lean_object* x_560; -x_560 = lean_ctor_get(x_558, 0); -lean_dec(x_560); -lean_ctor_set(x_558, 0, x_509); -return x_558; +return x_30; } else { -lean_object* x_561; lean_object* x_562; -x_561 = lean_ctor_get(x_558, 1); -lean_inc(x_561); -lean_dec(x_558); -x_562 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_562, 0, x_509); -lean_ctor_set(x_562, 1, x_561); -return x_562; +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_30, 0); +x_43 = lean_ctor_get(x_30, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_30); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } else { -lean_object* x_563; lean_object* x_564; lean_object* x_565; size_t x_566; uint64_t x_567; uint64_t x_568; uint64_t x_569; uint64_t x_570; uint64_t x_571; uint64_t x_572; uint64_t x_573; uint64_t x_574; uint64_t x_575; size_t x_576; size_t x_577; size_t x_578; size_t x_579; size_t x_580; lean_object* x_581; uint8_t x_582; -x_563 = lean_ctor_get(x_511, 0); -x_564 = lean_ctor_get(x_511, 1); -lean_inc(x_564); -lean_inc(x_563); -lean_dec(x_511); -x_565 = lean_array_get_size(x_564); -x_566 = lean_ptr_addr(x_1); -x_567 = lean_usize_to_uint64(x_566); -x_568 = 11; -x_569 = lean_uint64_mix_hash(x_567, x_568); -x_570 = 32; -x_571 = lean_uint64_shift_right(x_569, x_570); -x_572 = lean_uint64_xor(x_569, x_571); -x_573 = 16; -x_574 = lean_uint64_shift_right(x_572, x_573); -x_575 = lean_uint64_xor(x_572, x_574); -x_576 = lean_uint64_to_usize(x_575); -x_577 = lean_usize_of_nat(x_565); -lean_dec(x_565); -x_578 = 1; -x_579 = lean_usize_sub(x_577, x_578); -x_580 = lean_usize_land(x_576, x_579); -x_581 = lean_array_uget(x_564, x_580); -x_582 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_581); -if (x_582 == 0) -{ -lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; uint8_t x_591; -x_583 = lean_nat_add(x_563, x_501); -lean_dec(x_563); -lean_inc(x_509); -x_584 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_584, 0, x_1); -lean_ctor_set(x_584, 1, x_509); -lean_ctor_set(x_584, 2, x_581); -x_585 = lean_array_uset(x_564, x_580, x_584); -x_586 = lean_unsigned_to_nat(4u); -x_587 = lean_nat_mul(x_583, x_586); -x_588 = lean_unsigned_to_nat(3u); -x_589 = lean_nat_div(x_587, x_588); -lean_dec(x_587); -x_590 = lean_array_get_size(x_585); -x_591 = lean_nat_dec_le(x_589, x_590); -lean_dec(x_590); -lean_dec(x_589); -if (x_591 == 0) -{ -lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; -x_592 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_585); -x_593 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_593, 0, x_583); -lean_ctor_set(x_593, 1, x_592); -x_594 = lean_st_ref_set(x_5, x_593, x_512); -lean_dec(x_5); -x_595 = lean_ctor_get(x_594, 1); -lean_inc(x_595); -if (lean_is_exclusive(x_594)) { - lean_ctor_release(x_594, 0); - lean_ctor_release(x_594, 1); - x_596 = x_594; -} else { - lean_dec_ref(x_594); - x_596 = lean_box(0); -} -if (lean_is_scalar(x_596)) { - x_597 = lean_alloc_ctor(0, 2, 0); -} else { - x_597 = x_596; -} -lean_ctor_set(x_597, 0, x_509); -lean_ctor_set(x_597, 1, x_595); -return x_597; -} -else +lean_object* x_45; lean_object* x_46; +x_45 = lean_array_fget(x_24, x_4); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_45); +x_46 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_45, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_46) == 0) { -lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; -x_598 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_598, 0, x_583); -lean_ctor_set(x_598, 1, x_585); -x_599 = lean_st_ref_set(x_5, x_598, x_512); -lean_dec(x_5); -x_600 = lean_ctor_get(x_599, 1); -lean_inc(x_600); -if (lean_is_exclusive(x_599)) { - lean_ctor_release(x_599, 0); - lean_ctor_release(x_599, 1); - x_601 = x_599; -} else { - lean_dec_ref(x_599); - x_601 = lean_box(0); -} -if (lean_is_scalar(x_601)) { - x_602 = lean_alloc_ctor(0, 2, 0); -} else { - x_602 = x_601; -} -lean_ctor_set(x_602, 0, x_509); -lean_ctor_set(x_602, 1, x_600); -return x_602; -} +lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; uint8_t x_51; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_ptr_addr(x_45); +lean_dec(x_45); +x_50 = lean_ptr_addr(x_47); +x_51 = lean_usize_dec_eq(x_49, x_50); +if (x_51 == 0) +{ +lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_25); +x_52 = lean_array_set(x_24, x_4, x_47); +x_53 = 1; +x_54 = lean_box(x_53); +lean_ctor_set(x_3, 1, x_54); +lean_ctor_set(x_3, 0, x_52); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_3); +x_16 = x_55; +x_17 = x_48; +goto block_22; } else { -lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; -x_603 = lean_box(0); -x_604 = lean_array_uset(x_564, x_580, x_603); -lean_inc(x_509); -x_605 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_509, x_581); -x_606 = lean_array_uset(x_604, x_580, x_605); -x_607 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_607, 0, x_563); -lean_ctor_set(x_607, 1, x_606); -x_608 = lean_st_ref_set(x_5, x_607, x_512); -lean_dec(x_5); -x_609 = lean_ctor_get(x_608, 1); -lean_inc(x_609); -if (lean_is_exclusive(x_608)) { - lean_ctor_release(x_608, 0); - lean_ctor_release(x_608, 1); - x_610 = x_608; -} else { - lean_dec_ref(x_608); - x_610 = lean_box(0); -} -if (lean_is_scalar(x_610)) { - x_611 = lean_alloc_ctor(0, 2, 0); -} else { - x_611 = x_610; -} -lean_ctor_set(x_611, 0, x_509); -lean_ctor_set(x_611, 1, x_609); -return x_611; -} -} +lean_object* x_56; +lean_dec(x_47); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_3); +x_16 = x_56; +x_17 = x_48; +goto block_22; } } else { -uint8_t x_617; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_617 = !lean_is_exclusive(x_506); -if (x_617 == 0) +uint8_t x_57; +lean_dec(x_45); +lean_free_object(x_3); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_57 = !lean_is_exclusive(x_46); +if (x_57 == 0) { -return x_506; +return x_46; } else { -lean_object* x_618; lean_object* x_619; lean_object* x_620; -x_618 = lean_ctor_get(x_506, 0); -x_619 = lean_ctor_get(x_506, 1); -lean_inc(x_619); -lean_inc(x_618); -lean_dec(x_506); -x_620 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_620, 0, x_618); -lean_ctor_set(x_620, 1, x_619); -return x_620; -} +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_46, 0); +x_59 = lean_ctor_get(x_46, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_46); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } -case 5: -{ -lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; -x_621 = lean_ctor_get(x_2, 0); -lean_inc(x_621); -x_622 = lean_ctor_get(x_2, 1); -lean_inc(x_622); -lean_dec(x_2); -x_623 = lean_array_set(x_3, x_4, x_622); -x_624 = lean_unsigned_to_nat(1u); -x_625 = lean_nat_sub(x_4, x_624); -lean_dec(x_4); -x_2 = x_621; -x_3 = x_623; -x_4 = x_625; -goto _start; } -case 6: -{ -lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; uint8_t x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; -lean_dec(x_4); -x_627 = lean_array_get_size(x_3); -x_628 = lean_unsigned_to_nat(0u); -x_629 = lean_unsigned_to_nat(1u); -x_630 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_630, 0, x_628); -lean_ctor_set(x_630, 1, x_627); -lean_ctor_set(x_630, 2, x_629); -x_631 = 0; -x_632 = lean_box(x_631); -x_633 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_633, 0, x_3); -lean_ctor_set(x_633, 1, x_632); -lean_inc(x_5); -x_634 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_630, x_630, x_633, x_628, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_630); -if (lean_obj_tag(x_634) == 0) -{ -lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_741; uint8_t x_742; -x_635 = lean_ctor_get(x_634, 0); -lean_inc(x_635); -x_636 = lean_ctor_get(x_634, 1); -lean_inc(x_636); -lean_dec(x_634); -x_741 = lean_ctor_get(x_635, 1); -lean_inc(x_741); -x_742 = lean_unbox(x_741); -lean_dec(x_741); -if (x_742 == 0) -{ -lean_dec(x_635); -lean_dec(x_2); -lean_inc(x_1); -x_637 = x_1; -goto block_740; } else { -lean_object* x_743; lean_object* x_744; -x_743 = lean_ctor_get(x_635, 0); -lean_inc(x_743); -lean_dec(x_635); -x_744 = l_Lean_mkAppN(x_2, x_743); -lean_dec(x_743); -x_637 = x_744; -goto block_740; -} -block_740: -{ -lean_object* x_638; lean_object* x_639; lean_object* x_640; uint8_t x_641; -x_638 = lean_st_ref_take(x_5, x_636); -x_639 = lean_ctor_get(x_638, 0); -lean_inc(x_639); -x_640 = lean_ctor_get(x_638, 1); -lean_inc(x_640); -lean_dec(x_638); -x_641 = !lean_is_exclusive(x_639); -if (x_641 == 0) -{ -lean_object* x_642; lean_object* x_643; lean_object* x_644; size_t x_645; uint64_t x_646; uint64_t x_647; uint64_t x_648; uint64_t x_649; uint64_t x_650; uint64_t x_651; uint64_t x_652; uint64_t x_653; uint64_t x_654; size_t x_655; size_t x_656; size_t x_657; size_t x_658; size_t x_659; lean_object* x_660; uint8_t x_661; -x_642 = lean_ctor_get(x_639, 0); -x_643 = lean_ctor_get(x_639, 1); -x_644 = lean_array_get_size(x_643); -x_645 = lean_ptr_addr(x_1); -x_646 = lean_usize_to_uint64(x_645); -x_647 = 11; -x_648 = lean_uint64_mix_hash(x_646, x_647); -x_649 = 32; -x_650 = lean_uint64_shift_right(x_648, x_649); -x_651 = lean_uint64_xor(x_648, x_650); -x_652 = 16; -x_653 = lean_uint64_shift_right(x_651, x_652); -x_654 = lean_uint64_xor(x_651, x_653); -x_655 = lean_uint64_to_usize(x_654); -x_656 = lean_usize_of_nat(x_644); -lean_dec(x_644); -x_657 = 1; -x_658 = lean_usize_sub(x_656, x_657); -x_659 = lean_usize_land(x_655, x_658); -x_660 = lean_array_uget(x_643, x_659); -x_661 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_660); -if (x_661 == 0) -{ -lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; uint8_t x_670; -x_662 = lean_nat_add(x_642, x_629); -lean_dec(x_642); -lean_inc(x_637); -x_663 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_663, 0, x_1); -lean_ctor_set(x_663, 1, x_637); -lean_ctor_set(x_663, 2, x_660); -x_664 = lean_array_uset(x_643, x_659, x_663); -x_665 = lean_unsigned_to_nat(4u); -x_666 = lean_nat_mul(x_662, x_665); -x_667 = lean_unsigned_to_nat(3u); -x_668 = lean_nat_div(x_666, x_667); -lean_dec(x_666); -x_669 = lean_array_get_size(x_664); -x_670 = lean_nat_dec_le(x_668, x_669); -lean_dec(x_669); -lean_dec(x_668); -if (x_670 == 0) -{ -lean_object* x_671; lean_object* x_672; uint8_t x_673; -x_671 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_664); -lean_ctor_set(x_639, 1, x_671); -lean_ctor_set(x_639, 0, x_662); -x_672 = lean_st_ref_set(x_5, x_639, x_640); -lean_dec(x_5); -x_673 = !lean_is_exclusive(x_672); -if (x_673 == 0) +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_61 = lean_ctor_get(x_3, 0); +x_62 = lean_ctor_get(x_3, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_3); +x_63 = lean_array_get_size(x_61); +x_64 = lean_nat_dec_lt(x_4, x_63); +lean_dec(x_63); +if (x_64 == 0) +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = l_Lean_instInhabitedExpr; +x_66 = l_outOfBounds___rarg(x_65); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_66); +x_67 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_66, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_67) == 0) +{ +lean_object* x_68; lean_object* x_69; size_t x_70; size_t x_71; uint8_t x_72; +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = lean_ptr_addr(x_66); +lean_dec(x_66); +x_71 = lean_ptr_addr(x_68); +x_72 = lean_usize_dec_eq(x_70, x_71); +if (x_72 == 0) { -lean_object* x_674; -x_674 = lean_ctor_get(x_672, 0); -lean_dec(x_674); -lean_ctor_set(x_672, 0, x_637); -return x_672; +lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_62); +x_73 = lean_array_set(x_61, x_4, x_68); +x_74 = 1; +x_75 = lean_box(x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_73); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_76); +x_16 = x_77; +x_17 = x_69; +goto block_22; } else { -lean_object* x_675; lean_object* x_676; -x_675 = lean_ctor_get(x_672, 1); -lean_inc(x_675); -lean_dec(x_672); -x_676 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_676, 0, x_637); -lean_ctor_set(x_676, 1, x_675); -return x_676; +lean_object* x_78; lean_object* x_79; +lean_dec(x_68); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_61); +lean_ctor_set(x_78, 1, x_62); +x_79 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_79, 0, x_78); +x_16 = x_79; +x_17 = x_69; +goto block_22; } } else { -lean_object* x_677; uint8_t x_678; -lean_ctor_set(x_639, 1, x_664); -lean_ctor_set(x_639, 0, x_662); -x_677 = lean_st_ref_set(x_5, x_639, x_640); -lean_dec(x_5); -x_678 = !lean_is_exclusive(x_677); -if (x_678 == 0) -{ -lean_object* x_679; -x_679 = lean_ctor_get(x_677, 0); -lean_dec(x_679); -lean_ctor_set(x_677, 0, x_637); -return x_677; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_66); +lean_dec(x_62); +lean_dec(x_61); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_80 = lean_ctor_get(x_67, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_67, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_82 = x_67; +} else { + lean_dec_ref(x_67); + x_82 = lean_box(0); } -else -{ -lean_object* x_680; lean_object* x_681; -x_680 = lean_ctor_get(x_677, 1); -lean_inc(x_680); -lean_dec(x_677); -x_681 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_681, 0, x_637); -lean_ctor_set(x_681, 1, x_680); -return x_681; +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(1, 2, 0); +} else { + x_83 = x_82; } +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_81); +return x_83; } } else { -lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; uint8_t x_687; -x_682 = lean_box(0); -x_683 = lean_array_uset(x_643, x_659, x_682); -lean_inc(x_637); -x_684 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_637, x_660); -x_685 = lean_array_uset(x_683, x_659, x_684); -lean_ctor_set(x_639, 1, x_685); -x_686 = lean_st_ref_set(x_5, x_639, x_640); -lean_dec(x_5); -x_687 = !lean_is_exclusive(x_686); -if (x_687 == 0) +lean_object* x_84; lean_object* x_85; +x_84 = lean_array_fget(x_61, x_4); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_84); +x_85 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_84, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_688; -x_688 = lean_ctor_get(x_686, 0); -lean_dec(x_688); -lean_ctor_set(x_686, 0, x_637); -return x_686; -} -else +lean_object* x_86; lean_object* x_87; size_t x_88; size_t x_89; uint8_t x_90; +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); +lean_inc(x_87); +lean_dec(x_85); +x_88 = lean_ptr_addr(x_84); +lean_dec(x_84); +x_89 = lean_ptr_addr(x_86); +x_90 = lean_usize_dec_eq(x_88, x_89); +if (x_90 == 0) { -lean_object* x_689; lean_object* x_690; -x_689 = lean_ctor_get(x_686, 1); -lean_inc(x_689); -lean_dec(x_686); -x_690 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_690, 0, x_637); -lean_ctor_set(x_690, 1, x_689); -return x_690; -} -} +lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_62); +x_91 = lean_array_set(x_61, x_4, x_86); +x_92 = 1; +x_93 = lean_box(x_92); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_93); +x_95 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_95, 0, x_94); +x_16 = x_95; +x_17 = x_87; +goto block_22; } else { -lean_object* x_691; lean_object* x_692; lean_object* x_693; size_t x_694; uint64_t x_695; uint64_t x_696; uint64_t x_697; uint64_t x_698; uint64_t x_699; uint64_t x_700; uint64_t x_701; uint64_t x_702; uint64_t x_703; size_t x_704; size_t x_705; size_t x_706; size_t x_707; size_t x_708; lean_object* x_709; uint8_t x_710; -x_691 = lean_ctor_get(x_639, 0); -x_692 = lean_ctor_get(x_639, 1); -lean_inc(x_692); -lean_inc(x_691); -lean_dec(x_639); -x_693 = lean_array_get_size(x_692); -x_694 = lean_ptr_addr(x_1); -x_695 = lean_usize_to_uint64(x_694); -x_696 = 11; -x_697 = lean_uint64_mix_hash(x_695, x_696); -x_698 = 32; -x_699 = lean_uint64_shift_right(x_697, x_698); -x_700 = lean_uint64_xor(x_697, x_699); -x_701 = 16; -x_702 = lean_uint64_shift_right(x_700, x_701); -x_703 = lean_uint64_xor(x_700, x_702); -x_704 = lean_uint64_to_usize(x_703); -x_705 = lean_usize_of_nat(x_693); -lean_dec(x_693); -x_706 = 1; -x_707 = lean_usize_sub(x_705, x_706); -x_708 = lean_usize_land(x_704, x_707); -x_709 = lean_array_uget(x_692, x_708); -x_710 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_709); -if (x_710 == 0) -{ -lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; uint8_t x_719; -x_711 = lean_nat_add(x_691, x_629); -lean_dec(x_691); -lean_inc(x_637); -x_712 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_712, 0, x_1); -lean_ctor_set(x_712, 1, x_637); -lean_ctor_set(x_712, 2, x_709); -x_713 = lean_array_uset(x_692, x_708, x_712); -x_714 = lean_unsigned_to_nat(4u); -x_715 = lean_nat_mul(x_711, x_714); -x_716 = lean_unsigned_to_nat(3u); -x_717 = lean_nat_div(x_715, x_716); -lean_dec(x_715); -x_718 = lean_array_get_size(x_713); -x_719 = lean_nat_dec_le(x_717, x_718); -lean_dec(x_718); -lean_dec(x_717); -if (x_719 == 0) -{ -lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; -x_720 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_713); -x_721 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_721, 0, x_711); -lean_ctor_set(x_721, 1, x_720); -x_722 = lean_st_ref_set(x_5, x_721, x_640); -lean_dec(x_5); -x_723 = lean_ctor_get(x_722, 1); -lean_inc(x_723); -if (lean_is_exclusive(x_722)) { - lean_ctor_release(x_722, 0); - lean_ctor_release(x_722, 1); - x_724 = x_722; -} else { - lean_dec_ref(x_722); - x_724 = lean_box(0); -} -if (lean_is_scalar(x_724)) { - x_725 = lean_alloc_ctor(0, 2, 0); -} else { - x_725 = x_724; +lean_object* x_96; lean_object* x_97; +lean_dec(x_86); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_61); +lean_ctor_set(x_96, 1, x_62); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_16 = x_97; +x_17 = x_87; +goto block_22; } -lean_ctor_set(x_725, 0, x_637); -lean_ctor_set(x_725, 1, x_723); -return x_725; } else { -lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; -x_726 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_726, 0, x_711); -lean_ctor_set(x_726, 1, x_713); -x_727 = lean_st_ref_set(x_5, x_726, x_640); -lean_dec(x_5); -x_728 = lean_ctor_get(x_727, 1); -lean_inc(x_728); -if (lean_is_exclusive(x_727)) { - lean_ctor_release(x_727, 0); - lean_ctor_release(x_727, 1); - x_729 = x_727; +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +lean_dec(x_84); +lean_dec(x_62); +lean_dec(x_61); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +x_98 = lean_ctor_get(x_85, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_85, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_100 = x_85; } else { - lean_dec_ref(x_727); - x_729 = lean_box(0); + lean_dec_ref(x_85); + x_100 = lean_box(0); } -if (lean_is_scalar(x_729)) { - x_730 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_100)) { + x_101 = lean_alloc_ctor(1, 2, 0); } else { - x_730 = x_729; -} -lean_ctor_set(x_730, 0, x_637); -lean_ctor_set(x_730, 1, x_728); -return x_730; + x_101 = x_100; } +lean_ctor_set(x_101, 0, x_98); +lean_ctor_set(x_101, 1, x_99); +return x_101; } -else -{ -lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; -x_731 = lean_box(0); -x_732 = lean_array_uset(x_692, x_708, x_731); -lean_inc(x_637); -x_733 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_637, x_709); -x_734 = lean_array_uset(x_732, x_708, x_733); -x_735 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_735, 0, x_691); -lean_ctor_set(x_735, 1, x_734); -x_736 = lean_st_ref_set(x_5, x_735, x_640); -lean_dec(x_5); -x_737 = lean_ctor_get(x_736, 1); -lean_inc(x_737); -if (lean_is_exclusive(x_736)) { - lean_ctor_release(x_736, 0); - lean_ctor_release(x_736, 1); - x_738 = x_736; -} else { - lean_dec_ref(x_736); - x_738 = lean_box(0); } -if (lean_is_scalar(x_738)) { - x_739 = lean_alloc_ctor(0, 2, 0); -} else { - x_739 = x_738; } -lean_ctor_set(x_739, 0, x_637); -lean_ctor_set(x_739, 1, x_737); -return x_739; +block_22: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_2, 2); +x_20 = lean_nat_add(x_4, x_19); +lean_dec(x_4); +x_3 = x_18; +x_4 = x_20; +x_5 = lean_box(0); +x_6 = lean_box(0); +x_12 = x_17; +goto _start; } } } } -else +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_745; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_745 = !lean_is_exclusive(x_634); -if (x_745 == 0) +if (lean_obj_tag(x_2) == 5) { -return x_634; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +x_13 = lean_array_set(x_3, x_4, x_12); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_4, x_14); +lean_dec(x_4); +x_2 = x_11; +x_3 = x_13; +x_4 = x_15; +goto _start; } else { -lean_object* x_746; lean_object* x_747; lean_object* x_748; -x_746 = lean_ctor_get(x_634, 0); -x_747 = lean_ctor_get(x_634, 1); -lean_inc(x_747); -lean_inc(x_746); -lean_dec(x_634); -x_748 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_748, 0, x_746); -lean_ctor_set(x_748, 1, x_747); -return x_748; -} -} -} -case 7: -{ -lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; uint8_t x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_dec(x_4); -x_749 = lean_array_get_size(x_3); -x_750 = lean_unsigned_to_nat(0u); -x_751 = lean_unsigned_to_nat(1u); -x_752 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_752, 0, x_750); -lean_ctor_set(x_752, 1, x_749); -lean_ctor_set(x_752, 2, x_751); -x_753 = 0; -x_754 = lean_box(x_753); -x_755 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_755, 0, x_3); -lean_ctor_set(x_755, 1, x_754); -lean_inc(x_5); -x_756 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_752, x_752, x_755, x_750, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_752); -if (lean_obj_tag(x_756) == 0) -{ -lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_863; uint8_t x_864; -x_757 = lean_ctor_get(x_756, 0); -lean_inc(x_757); -x_758 = lean_ctor_get(x_756, 1); -lean_inc(x_758); -lean_dec(x_756); -x_863 = lean_ctor_get(x_757, 1); -lean_inc(x_863); -x_864 = lean_unbox(x_863); -lean_dec(x_863); -if (x_864 == 0) -{ -lean_dec(x_757); -lean_dec(x_2); -lean_inc(x_1); -x_759 = x_1; -goto block_862; -} -else +x_17 = lean_array_get_size(x_3); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_17); +lean_ctor_set(x_20, 2, x_19); +x_21 = 0; +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_3); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_20, x_20, x_23, x_18, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_20); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_865; lean_object* x_866; -x_865 = lean_ctor_get(x_757, 0); -lean_inc(x_865); -lean_dec(x_757); -x_866 = l_Lean_mkAppN(x_2, x_865); -lean_dec(x_865); -x_759 = x_866; -goto block_862; -} -block_862: -{ -lean_object* x_760; lean_object* x_761; lean_object* x_762; uint8_t x_763; -x_760 = lean_st_ref_take(x_5, x_758); -x_761 = lean_ctor_get(x_760, 0); -lean_inc(x_761); -x_762 = lean_ctor_get(x_760, 1); -lean_inc(x_762); -lean_dec(x_760); -x_763 = !lean_is_exclusive(x_761); -if (x_763 == 0) -{ -lean_object* x_764; lean_object* x_765; lean_object* x_766; size_t x_767; uint64_t x_768; uint64_t x_769; uint64_t x_770; uint64_t x_771; uint64_t x_772; uint64_t x_773; uint64_t x_774; uint64_t x_775; uint64_t x_776; size_t x_777; size_t x_778; size_t x_779; size_t x_780; size_t x_781; lean_object* x_782; uint8_t x_783; -x_764 = lean_ctor_get(x_761, 0); -x_765 = lean_ctor_get(x_761, 1); -x_766 = lean_array_get_size(x_765); -x_767 = lean_ptr_addr(x_1); -x_768 = lean_usize_to_uint64(x_767); -x_769 = 11; -x_770 = lean_uint64_mix_hash(x_768, x_769); -x_771 = 32; -x_772 = lean_uint64_shift_right(x_770, x_771); -x_773 = lean_uint64_xor(x_770, x_772); -x_774 = 16; -x_775 = lean_uint64_shift_right(x_773, x_774); -x_776 = lean_uint64_xor(x_773, x_775); -x_777 = lean_uint64_to_usize(x_776); -x_778 = lean_usize_of_nat(x_766); -lean_dec(x_766); -x_779 = 1; -x_780 = lean_usize_sub(x_778, x_779); -x_781 = lean_usize_land(x_777, x_780); -x_782 = lean_array_uget(x_765, x_781); -x_783 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_782); -if (x_783 == 0) -{ -lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; uint8_t x_792; -x_784 = lean_nat_add(x_764, x_751); -lean_dec(x_764); -lean_inc(x_759); -x_785 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_785, 0, x_1); -lean_ctor_set(x_785, 1, x_759); -lean_ctor_set(x_785, 2, x_782); -x_786 = lean_array_uset(x_765, x_781, x_785); -x_787 = lean_unsigned_to_nat(4u); -x_788 = lean_nat_mul(x_784, x_787); -x_789 = lean_unsigned_to_nat(3u); -x_790 = lean_nat_div(x_788, x_789); -lean_dec(x_788); -x_791 = lean_array_get_size(x_786); -x_792 = lean_nat_dec_le(x_790, x_791); -lean_dec(x_791); -lean_dec(x_790); -if (x_792 == 0) -{ -lean_object* x_793; lean_object* x_794; uint8_t x_795; -x_793 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_786); -lean_ctor_set(x_761, 1, x_793); -lean_ctor_set(x_761, 0, x_784); -x_794 = lean_st_ref_set(x_5, x_761, x_762); -lean_dec(x_5); -x_795 = !lean_is_exclusive(x_794); -if (x_795 == 0) +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +x_27 = lean_unbox(x_26); +lean_dec(x_26); +if (x_27 == 0) +{ +uint8_t x_28; +lean_dec(x_25); +lean_dec(x_2); +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -lean_object* x_796; -x_796 = lean_ctor_get(x_794, 0); -lean_dec(x_796); -lean_ctor_set(x_794, 0, x_759); -return x_794; +lean_object* x_29; +x_29 = lean_ctor_get(x_24, 0); +lean_dec(x_29); +lean_ctor_set(x_24, 0, x_1); +return x_24; } else { -lean_object* x_797; lean_object* x_798; -x_797 = lean_ctor_get(x_794, 1); -lean_inc(x_797); -lean_dec(x_794); -x_798 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_798, 0, x_759); -lean_ctor_set(x_798, 1, x_797); -return x_798; +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +lean_dec(x_24); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_1); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } else { -lean_object* x_799; uint8_t x_800; -lean_ctor_set(x_761, 1, x_786); -lean_ctor_set(x_761, 0, x_784); -x_799 = lean_st_ref_set(x_5, x_761, x_762); -lean_dec(x_5); -x_800 = !lean_is_exclusive(x_799); -if (x_800 == 0) +uint8_t x_32; +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_24); +if (x_32 == 0) { -lean_object* x_801; -x_801 = lean_ctor_get(x_799, 0); -lean_dec(x_801); -lean_ctor_set(x_799, 0, x_759); -return x_799; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_24, 0); +lean_dec(x_33); +x_34 = lean_ctor_get(x_25, 0); +lean_inc(x_34); +lean_dec(x_25); +x_35 = l_Lean_mkAppN(x_2, x_34); +lean_dec(x_34); +lean_ctor_set(x_24, 0, x_35); +return x_24; } else { -lean_object* x_802; lean_object* x_803; -x_802 = lean_ctor_get(x_799, 1); -lean_inc(x_802); -lean_dec(x_799); -x_803 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_803, 0, x_759); -lean_ctor_set(x_803, 1, x_802); -return x_803; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_24, 1); +lean_inc(x_36); +lean_dec(x_24); +x_37 = lean_ctor_get(x_25, 0); +lean_inc(x_37); +lean_dec(x_25); +x_38 = l_Lean_mkAppN(x_2, x_37); +lean_dec(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +return x_39; } } } else { -lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; uint8_t x_809; -x_804 = lean_box(0); -x_805 = lean_array_uset(x_765, x_781, x_804); -lean_inc(x_759); -x_806 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_759, x_782); -x_807 = lean_array_uset(x_805, x_781, x_806); -lean_ctor_set(x_761, 1, x_807); -x_808 = lean_st_ref_set(x_5, x_761, x_762); -lean_dec(x_5); -x_809 = !lean_is_exclusive(x_808); -if (x_809 == 0) +uint8_t x_40; +lean_dec(x_2); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_24); +if (x_40 == 0) { -lean_object* x_810; -x_810 = lean_ctor_get(x_808, 0); -lean_dec(x_810); -lean_ctor_set(x_808, 0, x_759); -return x_808; +return x_24; } else { -lean_object* x_811; lean_object* x_812; -x_811 = lean_ctor_get(x_808, 1); -lean_inc(x_811); -lean_dec(x_808); -x_812 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_812, 0, x_759); -lean_ctor_set(x_812, 1, x_811); -return x_812; -} +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_24, 0); +x_42 = lean_ctor_get(x_24, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_24); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } -else -{ -lean_object* x_813; lean_object* x_814; lean_object* x_815; size_t x_816; uint64_t x_817; uint64_t x_818; uint64_t x_819; uint64_t x_820; uint64_t x_821; uint64_t x_822; uint64_t x_823; uint64_t x_824; uint64_t x_825; size_t x_826; size_t x_827; size_t x_828; size_t x_829; size_t x_830; lean_object* x_831; uint8_t x_832; -x_813 = lean_ctor_get(x_761, 0); -x_814 = lean_ctor_get(x_761, 1); -lean_inc(x_814); -lean_inc(x_813); -lean_dec(x_761); -x_815 = lean_array_get_size(x_814); -x_816 = lean_ptr_addr(x_1); -x_817 = lean_usize_to_uint64(x_816); -x_818 = 11; -x_819 = lean_uint64_mix_hash(x_817, x_818); -x_820 = 32; -x_821 = lean_uint64_shift_right(x_819, x_820); -x_822 = lean_uint64_xor(x_819, x_821); -x_823 = 16; -x_824 = lean_uint64_shift_right(x_822, x_823); -x_825 = lean_uint64_xor(x_822, x_824); -x_826 = lean_uint64_to_usize(x_825); -x_827 = lean_usize_of_nat(x_815); -lean_dec(x_815); -x_828 = 1; -x_829 = lean_usize_sub(x_827, x_828); -x_830 = lean_usize_land(x_826, x_829); -x_831 = lean_array_uget(x_814, x_830); -x_832 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_831); -if (x_832 == 0) -{ -lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; uint8_t x_841; -x_833 = lean_nat_add(x_813, x_751); -lean_dec(x_813); -lean_inc(x_759); -x_834 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_834, 0, x_1); -lean_ctor_set(x_834, 1, x_759); -lean_ctor_set(x_834, 2, x_831); -x_835 = lean_array_uset(x_814, x_830, x_834); -x_836 = lean_unsigned_to_nat(4u); -x_837 = lean_nat_mul(x_833, x_836); -x_838 = lean_unsigned_to_nat(3u); -x_839 = lean_nat_div(x_837, x_838); -lean_dec(x_837); -x_840 = lean_array_get_size(x_835); -x_841 = lean_nat_dec_le(x_839, x_840); -lean_dec(x_840); -lean_dec(x_839); -if (x_841 == 0) -{ -lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; -x_842 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_835); -x_843 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_843, 0, x_833); -lean_ctor_set(x_843, 1, x_842); -x_844 = lean_st_ref_set(x_5, x_843, x_762); -lean_dec(x_5); -x_845 = lean_ctor_get(x_844, 1); -lean_inc(x_845); -if (lean_is_exclusive(x_844)) { - lean_ctor_release(x_844, 0); - lean_ctor_release(x_844, 1); - x_846 = x_844; -} else { - lean_dec_ref(x_844); - x_846 = lean_box(0); } -if (lean_is_scalar(x_846)) { - x_847 = lean_alloc_ctor(0, 2, 0); -} else { - x_847 = x_846; } -lean_ctor_set(x_847, 0, x_759); -lean_ctor_set(x_847, 1, x_845); -return x_847; } -else +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; -x_848 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_848, 0, x_833); -lean_ctor_set(x_848, 1, x_835); -x_849 = lean_st_ref_set(x_5, x_848, x_762); -lean_dec(x_5); -x_850 = lean_ctor_get(x_849, 1); -lean_inc(x_850); -if (lean_is_exclusive(x_849)) { - lean_ctor_release(x_849, 0); - lean_ctor_release(x_849, 1); - x_851 = x_849; -} else { - lean_dec_ref(x_849); - x_851 = lean_box(0); -} -if (lean_is_scalar(x_851)) { - x_852 = lean_alloc_ctor(0, 2, 0); -} else { - x_852 = x_851; -} -lean_ctor_set(x_852, 0, x_759); -lean_ctor_set(x_852, 1, x_850); -return x_852; -} -} -else +if (lean_obj_tag(x_2) == 0) { -lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; -x_853 = lean_box(0); -x_854 = lean_array_uset(x_814, x_830, x_853); -lean_inc(x_759); -x_855 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_759, x_831); -x_856 = lean_array_uset(x_854, x_830, x_855); -x_857 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_857, 0, x_813); -lean_ctor_set(x_857, 1, x_856); -x_858 = lean_st_ref_set(x_5, x_857, x_762); -lean_dec(x_5); -x_859 = lean_ctor_get(x_858, 1); -lean_inc(x_859); -if (lean_is_exclusive(x_858)) { - lean_ctor_release(x_858, 0); - lean_ctor_release(x_858, 1); - x_860 = x_858; -} else { - lean_dec_ref(x_858); - x_860 = lean_box(0); -} -if (lean_is_scalar(x_860)) { - x_861 = lean_alloc_ctor(0, 2, 0); -} else { - x_861 = x_860; -} -lean_ctor_set(x_861, 0, x_759); -lean_ctor_set(x_861, 1, x_859); -return x_861; -} -} -} +lean_object* x_3; +x_3 = lean_box(0); +return x_3; } else { -uint8_t x_867; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_867 = !lean_is_exclusive(x_756); -if (x_867 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = lean_ptr_addr(x_4); +x_8 = lean_ptr_addr(x_1); +x_9 = lean_usize_dec_eq(x_7, x_8); +if (x_9 == 0) { -return x_756; +x_2 = x_6; +goto _start; } else { -lean_object* x_868; lean_object* x_869; lean_object* x_870; -x_868 = lean_ctor_get(x_756, 0); -x_869 = lean_ctor_get(x_756, 1); -lean_inc(x_869); -lean_inc(x_868); -lean_dec(x_756); -x_870 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_870, 0, x_868); -lean_ctor_set(x_870, 1, x_869); -return x_870; +lean_object* x_11; +lean_inc(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_5); +return x_11; } } } -case 8: -{ -lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; uint8_t x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; -lean_dec(x_4); -x_871 = lean_array_get_size(x_3); -x_872 = lean_unsigned_to_nat(0u); -x_873 = lean_unsigned_to_nat(1u); -x_874 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_874, 0, x_872); -lean_ctor_set(x_874, 1, x_871); -lean_ctor_set(x_874, 2, x_873); -x_875 = 0; -x_876 = lean_box(x_875); -x_877 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_877, 0, x_3); -lean_ctor_set(x_877, 1, x_876); -lean_inc(x_5); -x_878 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_874, x_874, x_877, x_872, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_874); -if (lean_obj_tag(x_878) == 0) -{ -lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_985; uint8_t x_986; -x_879 = lean_ctor_get(x_878, 0); -lean_inc(x_879); -x_880 = lean_ctor_get(x_878, 1); -lean_inc(x_880); -lean_dec(x_878); -x_985 = lean_ctor_get(x_879, 1); -lean_inc(x_985); -x_986 = lean_unbox(x_985); -lean_dec(x_985); -if (x_986 == 0) -{ -lean_dec(x_879); -lean_dec(x_2); -lean_inc(x_1); -x_881 = x_1; -goto block_984; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_987; lean_object* x_988; -x_987 = lean_ctor_get(x_879, 0); -lean_inc(x_987); -lean_dec(x_879); -x_988 = l_Lean_mkAppN(x_2, x_987); -lean_dec(x_987); -x_881 = x_988; -goto block_984; -} -block_984: -{ -lean_object* x_882; lean_object* x_883; lean_object* x_884; uint8_t x_885; -x_882 = lean_st_ref_take(x_5, x_880); -x_883 = lean_ctor_get(x_882, 0); -lean_inc(x_883); -x_884 = lean_ctor_get(x_882, 1); -lean_inc(x_884); -lean_dec(x_882); -x_885 = !lean_is_exclusive(x_883); -if (x_885 == 0) -{ -lean_object* x_886; lean_object* x_887; lean_object* x_888; size_t x_889; uint64_t x_890; uint64_t x_891; uint64_t x_892; uint64_t x_893; uint64_t x_894; uint64_t x_895; uint64_t x_896; uint64_t x_897; uint64_t x_898; size_t x_899; size_t x_900; size_t x_901; size_t x_902; size_t x_903; lean_object* x_904; uint8_t x_905; -x_886 = lean_ctor_get(x_883, 0); -x_887 = lean_ctor_get(x_883, 1); -x_888 = lean_array_get_size(x_887); -x_889 = lean_ptr_addr(x_1); -x_890 = lean_usize_to_uint64(x_889); -x_891 = 11; -x_892 = lean_uint64_mix_hash(x_890, x_891); -x_893 = 32; -x_894 = lean_uint64_shift_right(x_892, x_893); -x_895 = lean_uint64_xor(x_892, x_894); -x_896 = 16; -x_897 = lean_uint64_shift_right(x_895, x_896); -x_898 = lean_uint64_xor(x_895, x_897); -x_899 = lean_uint64_to_usize(x_898); -x_900 = lean_usize_of_nat(x_888); -lean_dec(x_888); -x_901 = 1; -x_902 = lean_usize_sub(x_900, x_901); -x_903 = lean_usize_land(x_899, x_902); -x_904 = lean_array_uget(x_887, x_903); -x_905 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_904); -if (x_905 == 0) -{ -lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; uint8_t x_914; -x_906 = lean_nat_add(x_886, x_873); -lean_dec(x_886); -lean_inc(x_881); -x_907 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_907, 0, x_1); -lean_ctor_set(x_907, 1, x_881); -lean_ctor_set(x_907, 2, x_904); -x_908 = lean_array_uset(x_887, x_903, x_907); -x_909 = lean_unsigned_to_nat(4u); -x_910 = lean_nat_mul(x_906, x_909); -x_911 = lean_unsigned_to_nat(3u); -x_912 = lean_nat_div(x_910, x_911); -lean_dec(x_910); -x_913 = lean_array_get_size(x_908); -x_914 = lean_nat_dec_le(x_912, x_913); -lean_dec(x_913); -lean_dec(x_912); -if (x_914 == 0) -{ -lean_object* x_915; lean_object* x_916; uint8_t x_917; -x_915 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_908); -lean_ctor_set(x_883, 1, x_915); -lean_ctor_set(x_883, 0, x_906); -x_916 = lean_st_ref_set(x_5, x_883, x_884); -lean_dec(x_5); -x_917 = !lean_is_exclusive(x_916); -if (x_917 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_st_ref_take(x_3, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = !lean_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; size_t x_30; lean_object* x_31; uint8_t x_32; +x_13 = lean_ctor_get(x_10, 0); +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_array_get_size(x_14); +x_16 = lean_ptr_addr(x_1); +x_17 = lean_usize_to_uint64(x_16); +x_18 = 11; +x_19 = lean_uint64_mix_hash(x_17, x_18); +x_20 = 32; +x_21 = lean_uint64_shift_right(x_19, x_20); +x_22 = lean_uint64_xor(x_19, x_21); +x_23 = 16; +x_24 = lean_uint64_shift_right(x_22, x_23); +x_25 = lean_uint64_xor(x_22, x_24); +x_26 = lean_uint64_to_usize(x_25); +x_27 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_28 = 1; +x_29 = lean_usize_sub(x_27, x_28); +x_30 = lean_usize_land(x_26, x_29); +x_31 = lean_array_uget(x_14, x_30); +x_32 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_add(x_13, x_33); +lean_dec(x_13); +lean_inc(x_2); +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_2); +lean_ctor_set(x_35, 2, x_31); +x_36 = lean_array_uset(x_14, x_30, x_35); +x_37 = lean_unsigned_to_nat(4u); +x_38 = lean_nat_mul(x_34, x_37); +x_39 = lean_unsigned_to_nat(3u); +x_40 = lean_nat_div(x_38, x_39); +lean_dec(x_38); +x_41 = lean_array_get_size(x_36); +x_42 = lean_nat_dec_le(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_36); +lean_ctor_set(x_10, 1, x_43); +lean_ctor_set(x_10, 0, x_34); +x_44 = lean_st_ref_set(x_3, x_10, x_11); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) { -lean_object* x_918; -x_918 = lean_ctor_get(x_916, 0); -lean_dec(x_918); -lean_ctor_set(x_916, 0, x_881); -return x_916; +lean_object* x_46; +x_46 = lean_ctor_get(x_44, 0); +lean_dec(x_46); +lean_ctor_set(x_44, 0, x_2); +return x_44; } else { -lean_object* x_919; lean_object* x_920; -x_919 = lean_ctor_get(x_916, 1); -lean_inc(x_919); -lean_dec(x_916); -x_920 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_920, 0, x_881); -lean_ctor_set(x_920, 1, x_919); -return x_920; +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_2); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } else { -lean_object* x_921; uint8_t x_922; -lean_ctor_set(x_883, 1, x_908); -lean_ctor_set(x_883, 0, x_906); -x_921 = lean_st_ref_set(x_5, x_883, x_884); -lean_dec(x_5); -x_922 = !lean_is_exclusive(x_921); -if (x_922 == 0) +lean_object* x_49; uint8_t x_50; +lean_ctor_set(x_10, 1, x_36); +lean_ctor_set(x_10, 0, x_34); +x_49 = lean_st_ref_set(x_3, x_10, x_11); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) { -lean_object* x_923; -x_923 = lean_ctor_get(x_921, 0); -lean_dec(x_923); -lean_ctor_set(x_921, 0, x_881); -return x_921; +lean_object* x_51; +x_51 = lean_ctor_get(x_49, 0); +lean_dec(x_51); +lean_ctor_set(x_49, 0, x_2); +return x_49; } else { -lean_object* x_924; lean_object* x_925; -x_924 = lean_ctor_get(x_921, 1); -lean_inc(x_924); -lean_dec(x_921); -x_925 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_925, 0, x_881); -lean_ctor_set(x_925, 1, x_924); -return x_925; +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_49, 1); +lean_inc(x_52); +lean_dec(x_49); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_2); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } else { -lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; uint8_t x_931; -x_926 = lean_box(0); -x_927 = lean_array_uset(x_887, x_903, x_926); -lean_inc(x_881); -x_928 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_881, x_904); -x_929 = lean_array_uset(x_927, x_903, x_928); -lean_ctor_set(x_883, 1, x_929); -x_930 = lean_st_ref_set(x_5, x_883, x_884); -lean_dec(x_5); -x_931 = !lean_is_exclusive(x_930); -if (x_931 == 0) +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_54 = lean_box(0); +x_55 = lean_array_uset(x_14, x_30, x_54); +lean_inc(x_2); +x_56 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_2, x_31); +x_57 = lean_array_uset(x_55, x_30, x_56); +lean_ctor_set(x_10, 1, x_57); +x_58 = lean_st_ref_set(x_3, x_10, x_11); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) { -lean_object* x_932; -x_932 = lean_ctor_get(x_930, 0); -lean_dec(x_932); -lean_ctor_set(x_930, 0, x_881); -return x_930; +lean_object* x_60; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +lean_ctor_set(x_58, 0, x_2); +return x_58; } else { -lean_object* x_933; lean_object* x_934; -x_933 = lean_ctor_get(x_930, 1); -lean_inc(x_933); -lean_dec(x_930); -x_934 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_934, 0, x_881); -lean_ctor_set(x_934, 1, x_933); -return x_934; +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_2); +lean_ctor_set(x_62, 1, x_61); +return x_62; } } } else { -lean_object* x_935; lean_object* x_936; lean_object* x_937; size_t x_938; uint64_t x_939; uint64_t x_940; uint64_t x_941; uint64_t x_942; uint64_t x_943; uint64_t x_944; uint64_t x_945; uint64_t x_946; uint64_t x_947; size_t x_948; size_t x_949; size_t x_950; size_t x_951; size_t x_952; lean_object* x_953; uint8_t x_954; -x_935 = lean_ctor_get(x_883, 0); -x_936 = lean_ctor_get(x_883, 1); -lean_inc(x_936); -lean_inc(x_935); -lean_dec(x_883); -x_937 = lean_array_get_size(x_936); -x_938 = lean_ptr_addr(x_1); -x_939 = lean_usize_to_uint64(x_938); -x_940 = 11; -x_941 = lean_uint64_mix_hash(x_939, x_940); -x_942 = 32; -x_943 = lean_uint64_shift_right(x_941, x_942); -x_944 = lean_uint64_xor(x_941, x_943); -x_945 = 16; -x_946 = lean_uint64_shift_right(x_944, x_945); -x_947 = lean_uint64_xor(x_944, x_946); -x_948 = lean_uint64_to_usize(x_947); -x_949 = lean_usize_of_nat(x_937); -lean_dec(x_937); -x_950 = 1; -x_951 = lean_usize_sub(x_949, x_950); -x_952 = lean_usize_land(x_948, x_951); -x_953 = lean_array_uget(x_936, x_952); -x_954 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_953); -if (x_954 == 0) -{ -lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; uint8_t x_963; -x_955 = lean_nat_add(x_935, x_873); -lean_dec(x_935); -lean_inc(x_881); -x_956 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_956, 0, x_1); -lean_ctor_set(x_956, 1, x_881); -lean_ctor_set(x_956, 2, x_953); -x_957 = lean_array_uset(x_936, x_952, x_956); -x_958 = lean_unsigned_to_nat(4u); -x_959 = lean_nat_mul(x_955, x_958); -x_960 = lean_unsigned_to_nat(3u); -x_961 = lean_nat_div(x_959, x_960); -lean_dec(x_959); -x_962 = lean_array_get_size(x_957); -x_963 = lean_nat_dec_le(x_961, x_962); -lean_dec(x_962); -lean_dec(x_961); -if (x_963 == 0) -{ -lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; -x_964 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_957); -x_965 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_965, 0, x_955); -lean_ctor_set(x_965, 1, x_964); -x_966 = lean_st_ref_set(x_5, x_965, x_884); -lean_dec(x_5); -x_967 = lean_ctor_get(x_966, 1); -lean_inc(x_967); -if (lean_is_exclusive(x_966)) { - lean_ctor_release(x_966, 0); - lean_ctor_release(x_966, 1); - x_968 = x_966; +lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; uint64_t x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; size_t x_76; size_t x_77; size_t x_78; size_t x_79; size_t x_80; lean_object* x_81; uint8_t x_82; +x_63 = lean_ctor_get(x_10, 0); +x_64 = lean_ctor_get(x_10, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_10); +x_65 = lean_array_get_size(x_64); +x_66 = lean_ptr_addr(x_1); +x_67 = lean_usize_to_uint64(x_66); +x_68 = 11; +x_69 = lean_uint64_mix_hash(x_67, x_68); +x_70 = 32; +x_71 = lean_uint64_shift_right(x_69, x_70); +x_72 = lean_uint64_xor(x_69, x_71); +x_73 = 16; +x_74 = lean_uint64_shift_right(x_72, x_73); +x_75 = lean_uint64_xor(x_72, x_74); +x_76 = lean_uint64_to_usize(x_75); +x_77 = lean_usize_of_nat(x_65); +lean_dec(x_65); +x_78 = 1; +x_79 = lean_usize_sub(x_77, x_78); +x_80 = lean_usize_land(x_76, x_79); +x_81 = lean_array_uget(x_64, x_80); +x_82 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_81); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_83 = lean_unsigned_to_nat(1u); +x_84 = lean_nat_add(x_63, x_83); +lean_dec(x_63); +lean_inc(x_2); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_1); +lean_ctor_set(x_85, 1, x_2); +lean_ctor_set(x_85, 2, x_81); +x_86 = lean_array_uset(x_64, x_80, x_85); +x_87 = lean_unsigned_to_nat(4u); +x_88 = lean_nat_mul(x_84, x_87); +x_89 = lean_unsigned_to_nat(3u); +x_90 = lean_nat_div(x_88, x_89); +lean_dec(x_88); +x_91 = lean_array_get_size(x_86); +x_92 = lean_nat_dec_le(x_90, x_91); +lean_dec(x_91); +lean_dec(x_90); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_93 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_86); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_84); +lean_ctor_set(x_94, 1, x_93); +x_95 = lean_st_ref_set(x_3, x_94, x_11); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_97 = x_95; } else { - lean_dec_ref(x_966); - x_968 = lean_box(0); + lean_dec_ref(x_95); + x_97 = lean_box(0); } -if (lean_is_scalar(x_968)) { - x_969 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(0, 2, 0); } else { - x_969 = x_968; + x_98 = x_97; } -lean_ctor_set(x_969, 0, x_881); -lean_ctor_set(x_969, 1, x_967); -return x_969; +lean_ctor_set(x_98, 0, x_2); +lean_ctor_set(x_98, 1, x_96); +return x_98; } else { -lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; -x_970 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_970, 0, x_955); -lean_ctor_set(x_970, 1, x_957); -x_971 = lean_st_ref_set(x_5, x_970, x_884); -lean_dec(x_5); -x_972 = lean_ctor_get(x_971, 1); -lean_inc(x_972); -if (lean_is_exclusive(x_971)) { - lean_ctor_release(x_971, 0); - lean_ctor_release(x_971, 1); - x_973 = x_971; +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_84); +lean_ctor_set(x_99, 1, x_86); +x_100 = lean_st_ref_set(x_3, x_99, x_11); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_102 = x_100; } else { - lean_dec_ref(x_971); - x_973 = lean_box(0); + lean_dec_ref(x_100); + x_102 = lean_box(0); } -if (lean_is_scalar(x_973)) { - x_974 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_102)) { + x_103 = lean_alloc_ctor(0, 2, 0); } else { - x_974 = x_973; + x_103 = x_102; } -lean_ctor_set(x_974, 0, x_881); -lean_ctor_set(x_974, 1, x_972); -return x_974; +lean_ctor_set(x_103, 0, x_2); +lean_ctor_set(x_103, 1, x_101); +return x_103; } } else { -lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; -x_975 = lean_box(0); -x_976 = lean_array_uset(x_936, x_952, x_975); -lean_inc(x_881); -x_977 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_881, x_953); -x_978 = lean_array_uset(x_976, x_952, x_977); -x_979 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_979, 0, x_935); -lean_ctor_set(x_979, 1, x_978); -x_980 = lean_st_ref_set(x_5, x_979, x_884); -lean_dec(x_5); -x_981 = lean_ctor_get(x_980, 1); -lean_inc(x_981); -if (lean_is_exclusive(x_980)) { - lean_ctor_release(x_980, 0); - lean_ctor_release(x_980, 1); - x_982 = x_980; +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_104 = lean_box(0); +x_105 = lean_array_uset(x_64, x_80, x_104); +lean_inc(x_2); +x_106 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_2, x_81); +x_107 = lean_array_uset(x_105, x_80, x_106); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_63); +lean_ctor_set(x_108, 1, x_107); +x_109 = lean_st_ref_set(x_3, x_108, x_11); +x_110 = lean_ctor_get(x_109, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_111 = x_109; } else { - lean_dec_ref(x_980); - x_982 = lean_box(0); + lean_dec_ref(x_109); + x_111 = lean_box(0); } -if (lean_is_scalar(x_982)) { - x_983 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_111)) { + x_112 = lean_alloc_ctor(0, 2, 0); } else { - x_983 = x_982; + x_112 = x_111; } -lean_ctor_set(x_983, 0, x_881); -lean_ctor_set(x_983, 1, x_981); -return x_983; +lean_ctor_set(x_112, 0, x_2); +lean_ctor_set(x_112, 1, x_110); +return x_112; } } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: { -uint8_t x_989; -lean_dec(x_5); +lean_object* x_14; size_t x_15; size_t x_16; uint8_t x_17; +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l_Lean_Expr_forallE___override(x_1, x_2, x_3, x_4); +x_15 = lean_ptr_addr(x_2); lean_dec(x_2); -lean_dec(x_1); -x_989 = !lean_is_exclusive(x_878); -if (x_989 == 0) +x_16 = lean_ptr_addr(x_5); +x_17 = lean_usize_dec_eq(x_15, x_16); +if (x_17 == 0) { -return x_878; +lean_object* x_18; lean_object* x_19; +lean_dec(x_14); +lean_dec(x_3); +x_18 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_19 = lean_apply_7(x_6, x_18, x_8, x_9, x_10, x_11, x_12, x_13); +return x_19; } else { -lean_object* x_990; lean_object* x_991; lean_object* x_992; -x_990 = lean_ctor_get(x_878, 0); -x_991 = lean_ctor_get(x_878, 1); -lean_inc(x_991); -lean_inc(x_990); -lean_dec(x_878); -x_992 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_992, 0, x_990); -lean_ctor_set(x_992, 1, x_991); -return x_992; -} -} -} -case 9: +size_t x_20; size_t x_21; uint8_t x_22; +x_20 = lean_ptr_addr(x_3); +lean_dec(x_3); +x_21 = lean_ptr_addr(x_7); +x_22 = lean_usize_dec_eq(x_20, x_21); +if (x_22 == 0) { -lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; uint8_t x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; -lean_dec(x_4); -x_993 = lean_array_get_size(x_3); -x_994 = lean_unsigned_to_nat(0u); -x_995 = lean_unsigned_to_nat(1u); -x_996 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_996, 0, x_994); -lean_ctor_set(x_996, 1, x_993); -lean_ctor_set(x_996, 2, x_995); -x_997 = 0; -x_998 = lean_box(x_997); -x_999 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_999, 0, x_3); -lean_ctor_set(x_999, 1, x_998); -lean_inc(x_5); -x_1000 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_996, x_996, x_999, x_994, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_996); -if (lean_obj_tag(x_1000) == 0) -{ -lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1107; uint8_t x_1108; -x_1001 = lean_ctor_get(x_1000, 0); -lean_inc(x_1001); -x_1002 = lean_ctor_get(x_1000, 1); -lean_inc(x_1002); -lean_dec(x_1000); -x_1107 = lean_ctor_get(x_1001, 1); -lean_inc(x_1107); -x_1108 = lean_unbox(x_1107); -lean_dec(x_1107); -if (x_1108 == 0) -{ -lean_dec(x_1001); -lean_dec(x_2); -lean_inc(x_1); -x_1003 = x_1; -goto block_1106; +lean_object* x_23; lean_object* x_24; +lean_dec(x_14); +x_23 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_24 = lean_apply_7(x_6, x_23, x_8, x_9, x_10, x_11, x_12, x_13); +return x_24; } else { -lean_object* x_1109; lean_object* x_1110; -x_1109 = lean_ctor_get(x_1001, 0); -lean_inc(x_1109); -lean_dec(x_1001); -x_1110 = l_Lean_mkAppN(x_2, x_1109); -lean_dec(x_1109); -x_1003 = x_1110; -goto block_1106; -} -block_1106: -{ -lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; uint8_t x_1007; -x_1004 = lean_st_ref_take(x_5, x_1002); -x_1005 = lean_ctor_get(x_1004, 0); -lean_inc(x_1005); -x_1006 = lean_ctor_get(x_1004, 1); -lean_inc(x_1006); -lean_dec(x_1004); -x_1007 = !lean_is_exclusive(x_1005); -if (x_1007 == 0) -{ -lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; size_t x_1011; uint64_t x_1012; uint64_t x_1013; uint64_t x_1014; uint64_t x_1015; uint64_t x_1016; uint64_t x_1017; uint64_t x_1018; uint64_t x_1019; uint64_t x_1020; size_t x_1021; size_t x_1022; size_t x_1023; size_t x_1024; size_t x_1025; lean_object* x_1026; uint8_t x_1027; -x_1008 = lean_ctor_get(x_1005, 0); -x_1009 = lean_ctor_get(x_1005, 1); -x_1010 = lean_array_get_size(x_1009); -x_1011 = lean_ptr_addr(x_1); -x_1012 = lean_usize_to_uint64(x_1011); -x_1013 = 11; -x_1014 = lean_uint64_mix_hash(x_1012, x_1013); -x_1015 = 32; -x_1016 = lean_uint64_shift_right(x_1014, x_1015); -x_1017 = lean_uint64_xor(x_1014, x_1016); -x_1018 = 16; -x_1019 = lean_uint64_shift_right(x_1017, x_1018); -x_1020 = lean_uint64_xor(x_1017, x_1019); -x_1021 = lean_uint64_to_usize(x_1020); -x_1022 = lean_usize_of_nat(x_1010); -lean_dec(x_1010); -x_1023 = 1; -x_1024 = lean_usize_sub(x_1022, x_1023); -x_1025 = lean_usize_land(x_1021, x_1024); -x_1026 = lean_array_uget(x_1009, x_1025); -x_1027 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1026); -if (x_1027 == 0) -{ -lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; uint8_t x_1036; -x_1028 = lean_nat_add(x_1008, x_995); -lean_dec(x_1008); -lean_inc(x_1003); -x_1029 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1029, 0, x_1); -lean_ctor_set(x_1029, 1, x_1003); -lean_ctor_set(x_1029, 2, x_1026); -x_1030 = lean_array_uset(x_1009, x_1025, x_1029); -x_1031 = lean_unsigned_to_nat(4u); -x_1032 = lean_nat_mul(x_1028, x_1031); -x_1033 = lean_unsigned_to_nat(3u); -x_1034 = lean_nat_div(x_1032, x_1033); -lean_dec(x_1032); -x_1035 = lean_array_get_size(x_1030); -x_1036 = lean_nat_dec_le(x_1034, x_1035); -lean_dec(x_1035); -lean_dec(x_1034); -if (x_1036 == 0) -{ -lean_object* x_1037; lean_object* x_1038; uint8_t x_1039; -x_1037 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1030); -lean_ctor_set(x_1005, 1, x_1037); -lean_ctor_set(x_1005, 0, x_1028); -x_1038 = lean_st_ref_set(x_5, x_1005, x_1006); -lean_dec(x_5); -x_1039 = !lean_is_exclusive(x_1038); -if (x_1039 == 0) -{ -lean_object* x_1040; -x_1040 = lean_ctor_get(x_1038, 0); -lean_dec(x_1040); -lean_ctor_set(x_1038, 0, x_1003); -return x_1038; -} -else +uint8_t x_25; +x_25 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(x_4, x_4); +if (x_25 == 0) { -lean_object* x_1041; lean_object* x_1042; -x_1041 = lean_ctor_get(x_1038, 1); -lean_inc(x_1041); -lean_dec(x_1038); -x_1042 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1042, 0, x_1003); -lean_ctor_set(x_1042, 1, x_1041); -return x_1042; -} +lean_object* x_26; lean_object* x_27; +lean_dec(x_14); +x_26 = l_Lean_Expr_forallE___override(x_1, x_5, x_7, x_4); +x_27 = lean_apply_7(x_6, x_26, x_8, x_9, x_10, x_11, x_12, x_13); +return x_27; } else { -lean_object* x_1043; uint8_t x_1044; -lean_ctor_set(x_1005, 1, x_1030); -lean_ctor_set(x_1005, 0, x_1028); -x_1043 = lean_st_ref_set(x_5, x_1005, x_1006); +lean_object* x_28; +lean_dec(x_7); lean_dec(x_5); -x_1044 = !lean_is_exclusive(x_1043); -if (x_1044 == 0) -{ -lean_object* x_1045; -x_1045 = lean_ctor_get(x_1043, 0); -lean_dec(x_1045); -lean_ctor_set(x_1043, 0, x_1003); -return x_1043; +lean_dec(x_1); +x_28 = lean_apply_7(x_6, x_14, x_8, x_9, x_10, x_11, x_12, x_13); +return x_28; } -else -{ -lean_object* x_1046; lean_object* x_1047; -x_1046 = lean_ctor_get(x_1043, 1); -lean_inc(x_1046); -lean_dec(x_1043); -x_1047 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1047, 0, x_1003); -lean_ctor_set(x_1047, 1, x_1046); -return x_1047; } } } -else -{ -lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; uint8_t x_1053; -x_1048 = lean_box(0); -x_1049 = lean_array_uset(x_1009, x_1025, x_1048); -lean_inc(x_1003); -x_1050 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1003, x_1026); -x_1051 = lean_array_uset(x_1049, x_1025, x_1050); -lean_ctor_set(x_1005, 1, x_1051); -x_1052 = lean_st_ref_set(x_5, x_1005, x_1006); -lean_dec(x_5); -x_1053 = !lean_is_exclusive(x_1052); -if (x_1053 == 0) -{ -lean_object* x_1054; -x_1054 = lean_ctor_get(x_1052, 0); -lean_dec(x_1054); -lean_ctor_set(x_1052, 0, x_1003); -return x_1052; } -else +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1() { +_start: { -lean_object* x_1055; lean_object* x_1056; -x_1055 = lean_ctor_get(x_1052, 1); -lean_inc(x_1055); -lean_dec(x_1052); -x_1056 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1056, 0, x_1003); -lean_ctor_set(x_1056, 1, x_1055); -return x_1056; -} +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.MarkNestedProofs", 39, 39); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2() { +_start: { -lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; size_t x_1060; uint64_t x_1061; uint64_t x_1062; uint64_t x_1063; uint64_t x_1064; uint64_t x_1065; uint64_t x_1066; uint64_t x_1067; uint64_t x_1068; uint64_t x_1069; size_t x_1070; size_t x_1071; size_t x_1072; size_t x_1073; size_t x_1074; lean_object* x_1075; uint8_t x_1076; -x_1057 = lean_ctor_get(x_1005, 0); -x_1058 = lean_ctor_get(x_1005, 1); -lean_inc(x_1058); -lean_inc(x_1057); -lean_dec(x_1005); -x_1059 = lean_array_get_size(x_1058); -x_1060 = lean_ptr_addr(x_1); -x_1061 = lean_usize_to_uint64(x_1060); -x_1062 = 11; -x_1063 = lean_uint64_mix_hash(x_1061, x_1062); -x_1064 = 32; -x_1065 = lean_uint64_shift_right(x_1063, x_1064); -x_1066 = lean_uint64_xor(x_1063, x_1065); -x_1067 = 16; -x_1068 = lean_uint64_shift_right(x_1066, x_1067); -x_1069 = lean_uint64_xor(x_1066, x_1068); -x_1070 = lean_uint64_to_usize(x_1069); -x_1071 = lean_usize_of_nat(x_1059); -lean_dec(x_1059); -x_1072 = 1; -x_1073 = lean_usize_sub(x_1071, x_1072); -x_1074 = lean_usize_land(x_1070, x_1073); -x_1075 = lean_array_uget(x_1058, x_1074); -x_1076 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1075); -if (x_1076 == 0) -{ -lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; uint8_t x_1085; -x_1077 = lean_nat_add(x_1057, x_995); -lean_dec(x_1057); -lean_inc(x_1003); -x_1078 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1078, 0, x_1); -lean_ctor_set(x_1078, 1, x_1003); -lean_ctor_set(x_1078, 2, x_1075); -x_1079 = lean_array_uset(x_1058, x_1074, x_1078); -x_1080 = lean_unsigned_to_nat(4u); -x_1081 = lean_nat_mul(x_1077, x_1080); -x_1082 = lean_unsigned_to_nat(3u); -x_1083 = lean_nat_div(x_1081, x_1082); -lean_dec(x_1081); -x_1084 = lean_array_get_size(x_1079); -x_1085 = lean_nat_dec_le(x_1083, x_1084); -lean_dec(x_1084); -lean_dec(x_1083); -if (x_1085 == 0) -{ -lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; -x_1086 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1079); -x_1087 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1087, 0, x_1077); -lean_ctor_set(x_1087, 1, x_1086); -x_1088 = lean_st_ref_set(x_5, x_1087, x_1006); -lean_dec(x_5); -x_1089 = lean_ctor_get(x_1088, 1); -lean_inc(x_1089); -if (lean_is_exclusive(x_1088)) { - lean_ctor_release(x_1088, 0); - lean_ctor_release(x_1088, 1); - x_1090 = x_1088; -} else { - lean_dec_ref(x_1088); - x_1090 = lean_box(0); -} -if (lean_is_scalar(x_1090)) { - x_1091 = lean_alloc_ctor(0, 2, 0); -} else { - x_1091 = x_1090; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.markNestedProofsImpl.visit", 42, 42); +return x_1; } -lean_ctor_set(x_1091, 0, x_1003); -lean_ctor_set(x_1091, 1, x_1089); -return x_1091; } -else +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3() { +_start: { -lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; -x_1092 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1092, 0, x_1077); -lean_ctor_set(x_1092, 1, x_1079); -x_1093 = lean_st_ref_set(x_5, x_1092, x_1006); -lean_dec(x_5); -x_1094 = lean_ctor_get(x_1093, 1); -lean_inc(x_1094); -if (lean_is_exclusive(x_1093)) { - lean_ctor_release(x_1093, 0); - lean_ctor_release(x_1093, 1); - x_1095 = x_1093; -} else { - lean_dec_ref(x_1093); - x_1095 = lean_box(0); -} -if (lean_is_scalar(x_1095)) { - x_1096 = lean_alloc_ctor(0, 2, 0); -} else { - x_1096 = x_1095; -} -lean_ctor_set(x_1096, 0, x_1003); -lean_ctor_set(x_1096, 1, x_1094); -return x_1096; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4() { +_start: { -lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; -x_1097 = lean_box(0); -x_1098 = lean_array_uset(x_1058, x_1074, x_1097); -lean_inc(x_1003); -x_1099 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1003, x_1075); -x_1100 = lean_array_uset(x_1098, x_1074, x_1099); -x_1101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1101, 0, x_1057); -lean_ctor_set(x_1101, 1, x_1100); -x_1102 = lean_st_ref_set(x_5, x_1101, x_1006); -lean_dec(x_5); -x_1103 = lean_ctor_get(x_1102, 1); -lean_inc(x_1103); -if (lean_is_exclusive(x_1102)) { - lean_ctor_release(x_1102, 0); - lean_ctor_release(x_1102, 1); - x_1104 = x_1102; -} else { - lean_dec_ref(x_1102); - x_1104 = lean_box(0); -} -if (lean_is_scalar(x_1104)) { - x_1105 = lean_alloc_ctor(0, 2, 0); -} else { - x_1105 = x_1104; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1; +x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2; +x_3 = lean_unsigned_to_nat(55u); +x_4 = lean_unsigned_to_nat(13u); +x_5 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -lean_ctor_set(x_1105, 0, x_1003); -lean_ctor_set(x_1105, 1, x_1103); -return x_1105; } +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 5: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_9); +x_11 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5; +lean_inc(x_10); +x_12 = lean_mk_array(x_10, x_11); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_sub(x_10, x_13); +lean_dec(x_10); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc_n(x_1, 2); +x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(x_1, x_1, x_12, x_14, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_18; } else { -uint8_t x_1111; +uint8_t x_19; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_1111 = !lean_is_exclusive(x_1000); -if (x_1111 == 0) +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) { -return x_1000; +return x_15; } else { -lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; -x_1112 = lean_ctor_get(x_1000, 0); -x_1113 = lean_ctor_get(x_1000, 1); -lean_inc(x_1113); -lean_inc(x_1112); -lean_dec(x_1000); -x_1114 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1114, 0, x_1112); -lean_ctor_set(x_1114, 1, x_1113); -return x_1114; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } -case 10: +case 7: { -lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; uint8_t x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; -lean_dec(x_4); -x_1115 = lean_array_get_size(x_3); -x_1116 = lean_unsigned_to_nat(0u); -x_1117 = lean_unsigned_to_nat(1u); -x_1118 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_1118, 0, x_1116); -lean_ctor_set(x_1118, 1, x_1115); -lean_ctor_set(x_1118, 2, x_1117); -x_1119 = 0; -x_1120 = lean_box(x_1119); -x_1121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1121, 0, x_3); -lean_ctor_set(x_1121, 1, x_1120); +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; +x_23 = lean_ctor_get(x_1, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_1, 2); +lean_inc(x_25); +x_26 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +x_27 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___boxed), 8, 1); +lean_closure_set(x_27, 0, x_1); +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -x_1122 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_1118, x_1118, x_1121, x_1116, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1118); -if (lean_obj_tag(x_1122) == 0) -{ -lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1229; uint8_t x_1230; -x_1123 = lean_ctor_get(x_1122, 0); -lean_inc(x_1123); -x_1124 = lean_ctor_get(x_1122, 1); -lean_inc(x_1124); -lean_dec(x_1122); -x_1229 = lean_ctor_get(x_1123, 1); -lean_inc(x_1229); -x_1230 = lean_unbox(x_1229); -lean_dec(x_1229); -if (x_1230 == 0) -{ -lean_dec(x_1123); -lean_dec(x_2); -lean_inc(x_1); -x_1125 = x_1; -goto block_1228; +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_24); +x_28 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_24, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Expr_hasLooseBVars(x_25); +if (x_31 == 0) +{ +lean_object* x_32; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_25); +x_32 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_25, x_3, x_4, x_5, x_6, x_7, x_30); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_23, x_24, x_25, x_26, x_29, x_27, x_33, x_3, x_4, x_5, x_6, x_7, x_34); +return x_35; } else { -lean_object* x_1231; lean_object* x_1232; -x_1231 = lean_ctor_get(x_1123, 0); -lean_inc(x_1231); -lean_dec(x_1123); -x_1232 = l_Lean_mkAppN(x_2, x_1231); -lean_dec(x_1231); -x_1125 = x_1232; -goto block_1228; -} -block_1228: -{ -lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; uint8_t x_1129; -x_1126 = lean_st_ref_take(x_5, x_1124); -x_1127 = lean_ctor_get(x_1126, 0); -lean_inc(x_1127); -x_1128 = lean_ctor_get(x_1126, 1); -lean_inc(x_1128); -lean_dec(x_1126); -x_1129 = !lean_is_exclusive(x_1127); -if (x_1129 == 0) -{ -lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; size_t x_1133; uint64_t x_1134; uint64_t x_1135; uint64_t x_1136; uint64_t x_1137; uint64_t x_1138; uint64_t x_1139; uint64_t x_1140; uint64_t x_1141; uint64_t x_1142; size_t x_1143; size_t x_1144; size_t x_1145; size_t x_1146; size_t x_1147; lean_object* x_1148; uint8_t x_1149; -x_1130 = lean_ctor_get(x_1127, 0); -x_1131 = lean_ctor_get(x_1127, 1); -x_1132 = lean_array_get_size(x_1131); -x_1133 = lean_ptr_addr(x_1); -x_1134 = lean_usize_to_uint64(x_1133); -x_1135 = 11; -x_1136 = lean_uint64_mix_hash(x_1134, x_1135); -x_1137 = 32; -x_1138 = lean_uint64_shift_right(x_1136, x_1137); -x_1139 = lean_uint64_xor(x_1136, x_1138); -x_1140 = 16; -x_1141 = lean_uint64_shift_right(x_1139, x_1140); -x_1142 = lean_uint64_xor(x_1139, x_1141); -x_1143 = lean_uint64_to_usize(x_1142); -x_1144 = lean_usize_of_nat(x_1132); -lean_dec(x_1132); -x_1145 = 1; -x_1146 = lean_usize_sub(x_1144, x_1145); -x_1147 = lean_usize_land(x_1143, x_1146); -x_1148 = lean_array_uget(x_1131, x_1147); -x_1149 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1148); -if (x_1149 == 0) -{ -lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; uint8_t x_1158; -x_1150 = lean_nat_add(x_1130, x_1117); -lean_dec(x_1130); -lean_inc(x_1125); -x_1151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1151, 0, x_1); -lean_ctor_set(x_1151, 1, x_1125); -lean_ctor_set(x_1151, 2, x_1148); -x_1152 = lean_array_uset(x_1131, x_1147, x_1151); -x_1153 = lean_unsigned_to_nat(4u); -x_1154 = lean_nat_mul(x_1150, x_1153); -x_1155 = lean_unsigned_to_nat(3u); -x_1156 = lean_nat_div(x_1154, x_1155); -lean_dec(x_1154); -x_1157 = lean_array_get_size(x_1152); -x_1158 = lean_nat_dec_le(x_1156, x_1157); -lean_dec(x_1157); -lean_dec(x_1156); -if (x_1158 == 0) -{ -lean_object* x_1159; lean_object* x_1160; uint8_t x_1161; -x_1159 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1152); -lean_ctor_set(x_1127, 1, x_1159); -lean_ctor_set(x_1127, 0, x_1150); -x_1160 = lean_st_ref_set(x_5, x_1127, x_1128); +uint8_t x_36; +lean_dec(x_29); +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1161 = !lean_is_exclusive(x_1160); -if (x_1161 == 0) +lean_dec(x_4); +lean_dec(x_3); +x_36 = !lean_is_exclusive(x_32); +if (x_36 == 0) { -lean_object* x_1162; -x_1162 = lean_ctor_get(x_1160, 0); -lean_dec(x_1162); -lean_ctor_set(x_1160, 0, x_1125); -return x_1160; +return x_32; } else { -lean_object* x_1163; lean_object* x_1164; -x_1163 = lean_ctor_get(x_1160, 1); -lean_inc(x_1163); -lean_dec(x_1160); -x_1164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1164, 0, x_1125); -lean_ctor_set(x_1164, 1, x_1163); -return x_1164; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_32, 0); +x_38 = lean_ctor_get(x_32, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_32); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } -else -{ -lean_object* x_1165; uint8_t x_1166; -lean_ctor_set(x_1127, 1, x_1152); -lean_ctor_set(x_1127, 0, x_1150); -x_1165 = lean_st_ref_set(x_5, x_1127, x_1128); -lean_dec(x_5); -x_1166 = !lean_is_exclusive(x_1165); -if (x_1166 == 0) -{ -lean_object* x_1167; -x_1167 = lean_ctor_get(x_1165, 0); -lean_dec(x_1167); -lean_ctor_set(x_1165, 0, x_1125); -return x_1165; } else { -lean_object* x_1168; lean_object* x_1169; -x_1168 = lean_ctor_get(x_1165, 1); -lean_inc(x_1168); -lean_dec(x_1165); -x_1169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1169, 0, x_1125); -lean_ctor_set(x_1169, 1, x_1168); -return x_1169; -} +lean_object* x_40; +lean_inc(x_25); +x_40 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_23, x_24, x_25, x_26, x_29, x_27, x_25, x_3, x_4, x_5, x_6, x_7, x_30); +return x_40; } } else { -lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; uint8_t x_1175; -x_1170 = lean_box(0); -x_1171 = lean_array_uset(x_1131, x_1147, x_1170); -lean_inc(x_1125); -x_1172 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1125, x_1148); -x_1173 = lean_array_uset(x_1171, x_1147, x_1172); -lean_ctor_set(x_1127, 1, x_1173); -x_1174 = lean_st_ref_set(x_5, x_1127, x_1128); +uint8_t x_41; +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1175 = !lean_is_exclusive(x_1174); -if (x_1175 == 0) +lean_dec(x_4); +lean_dec(x_3); +x_41 = !lean_is_exclusive(x_28); +if (x_41 == 0) { -lean_object* x_1176; -x_1176 = lean_ctor_get(x_1174, 0); -lean_dec(x_1176); -lean_ctor_set(x_1174, 0, x_1125); -return x_1174; +return x_28; } else { -lean_object* x_1177; lean_object* x_1178; -x_1177 = lean_ctor_get(x_1174, 1); -lean_inc(x_1177); -lean_dec(x_1174); -x_1178 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1178, 0, x_1125); -lean_ctor_set(x_1178, 1, x_1177); -return x_1178; +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_28, 0); +x_43 = lean_ctor_get(x_28, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_28); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } -else +case 11: { -lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; size_t x_1182; uint64_t x_1183; uint64_t x_1184; uint64_t x_1185; uint64_t x_1186; uint64_t x_1187; uint64_t x_1188; uint64_t x_1189; uint64_t x_1190; uint64_t x_1191; size_t x_1192; size_t x_1193; size_t x_1194; size_t x_1195; size_t x_1196; lean_object* x_1197; uint8_t x_1198; -x_1179 = lean_ctor_get(x_1127, 0); -x_1180 = lean_ctor_get(x_1127, 1); -lean_inc(x_1180); -lean_inc(x_1179); -lean_dec(x_1127); -x_1181 = lean_array_get_size(x_1180); -x_1182 = lean_ptr_addr(x_1); -x_1183 = lean_usize_to_uint64(x_1182); -x_1184 = 11; -x_1185 = lean_uint64_mix_hash(x_1183, x_1184); -x_1186 = 32; -x_1187 = lean_uint64_shift_right(x_1185, x_1186); -x_1188 = lean_uint64_xor(x_1185, x_1187); -x_1189 = 16; -x_1190 = lean_uint64_shift_right(x_1188, x_1189); -x_1191 = lean_uint64_xor(x_1188, x_1190); -x_1192 = lean_uint64_to_usize(x_1191); -x_1193 = lean_usize_of_nat(x_1181); -lean_dec(x_1181); -x_1194 = 1; -x_1195 = lean_usize_sub(x_1193, x_1194); -x_1196 = lean_usize_land(x_1192, x_1195); -x_1197 = lean_array_uget(x_1180, x_1196); -x_1198 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1197); -if (x_1198 == 0) -{ -lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; uint8_t x_1207; -x_1199 = lean_nat_add(x_1179, x_1117); -lean_dec(x_1179); -lean_inc(x_1125); -x_1200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1200, 0, x_1); -lean_ctor_set(x_1200, 1, x_1125); -lean_ctor_set(x_1200, 2, x_1197); -x_1201 = lean_array_uset(x_1180, x_1196, x_1200); -x_1202 = lean_unsigned_to_nat(4u); -x_1203 = lean_nat_mul(x_1199, x_1202); -x_1204 = lean_unsigned_to_nat(3u); -x_1205 = lean_nat_div(x_1203, x_1204); -lean_dec(x_1203); -x_1206 = lean_array_get_size(x_1201); -x_1207 = lean_nat_dec_le(x_1205, x_1206); -lean_dec(x_1206); -lean_dec(x_1205); -if (x_1207 == 0) -{ -lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; -x_1208 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1201); -x_1209 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1209, 0, x_1199); -lean_ctor_set(x_1209, 1, x_1208); -x_1210 = lean_st_ref_set(x_5, x_1209, x_1128); -lean_dec(x_5); -x_1211 = lean_ctor_get(x_1210, 1); -lean_inc(x_1211); -if (lean_is_exclusive(x_1210)) { - lean_ctor_release(x_1210, 0); - lean_ctor_release(x_1210, 1); - x_1212 = x_1210; -} else { - lean_dec_ref(x_1210); - x_1212 = lean_box(0); -} -if (lean_is_scalar(x_1212)) { - x_1213 = lean_alloc_ctor(0, 2, 0); -} else { - x_1213 = x_1212; -} -lean_ctor_set(x_1213, 0, x_1125); -lean_ctor_set(x_1213, 1, x_1211); -return x_1213; -} -else +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_1, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_1, 1); +lean_inc(x_46); +x_47 = lean_ctor_get(x_1, 2); +lean_inc(x_47); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_47); +x_48 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_47, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; size_t x_51; size_t x_52; uint8_t x_53; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = lean_ptr_addr(x_47); +lean_dec(x_47); +x_52 = lean_ptr_addr(x_49); +x_53 = lean_usize_dec_eq(x_51, x_52); +if (x_53 == 0) { -lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; -x_1214 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1214, 0, x_1199); -lean_ctor_set(x_1214, 1, x_1201); -x_1215 = lean_st_ref_set(x_5, x_1214, x_1128); +lean_object* x_54; lean_object* x_55; +x_54 = l_Lean_Expr_proj___override(x_45, x_46, x_49); +x_55 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_54, x_3, x_4, x_5, x_6, x_7, x_50); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1216 = lean_ctor_get(x_1215, 1); -lean_inc(x_1216); -if (lean_is_exclusive(x_1215)) { - lean_ctor_release(x_1215, 0); - lean_ctor_release(x_1215, 1); - x_1217 = x_1215; -} else { - lean_dec_ref(x_1215); - x_1217 = lean_box(0); -} -if (lean_is_scalar(x_1217)) { - x_1218 = lean_alloc_ctor(0, 2, 0); -} else { - x_1218 = x_1217; -} -lean_ctor_set(x_1218, 0, x_1125); -lean_ctor_set(x_1218, 1, x_1216); -return x_1218; -} +lean_dec(x_4); +lean_dec(x_3); +return x_55; } else { -lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; -x_1219 = lean_box(0); -x_1220 = lean_array_uset(x_1180, x_1196, x_1219); -lean_inc(x_1125); -x_1221 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1125, x_1197); -x_1222 = lean_array_uset(x_1220, x_1196, x_1221); -x_1223 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1223, 0, x_1179); -lean_ctor_set(x_1223, 1, x_1222); -x_1224 = lean_st_ref_set(x_5, x_1223, x_1128); +lean_object* x_56; +lean_dec(x_49); +lean_dec(x_46); +lean_dec(x_45); +lean_inc(x_1); +x_56 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_1, x_3, x_4, x_5, x_6, x_7, x_50); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1225 = lean_ctor_get(x_1224, 1); -lean_inc(x_1225); -if (lean_is_exclusive(x_1224)) { - lean_ctor_release(x_1224, 0); - lean_ctor_release(x_1224, 1); - x_1226 = x_1224; -} else { - lean_dec_ref(x_1224); - x_1226 = lean_box(0); -} -if (lean_is_scalar(x_1226)) { - x_1227 = lean_alloc_ctor(0, 2, 0); -} else { - x_1227 = x_1226; -} -lean_ctor_set(x_1227, 0, x_1125); -lean_ctor_set(x_1227, 1, x_1225); -return x_1227; -} -} +lean_dec(x_4); +lean_dec(x_3); +return x_56; } } else { -uint8_t x_1233; +uint8_t x_57; +lean_dec(x_47); +lean_dec(x_46); +lean_dec(x_45); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_1233 = !lean_is_exclusive(x_1122); -if (x_1233 == 0) +x_57 = !lean_is_exclusive(x_48); +if (x_57 == 0) { -return x_1122; +return x_48; } else { -lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; -x_1234 = lean_ctor_get(x_1122, 0); -x_1235 = lean_ctor_get(x_1122, 1); -lean_inc(x_1235); -lean_inc(x_1234); -lean_dec(x_1122); -x_1236 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1236, 0, x_1234); -lean_ctor_set(x_1236, 1, x_1235); -return x_1236; +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_48, 0); +x_59 = lean_ctor_get(x_48, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_48); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } } default: { -lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; uint8_t x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; -lean_dec(x_4); -x_1237 = lean_array_get_size(x_3); -x_1238 = lean_unsigned_to_nat(0u); -x_1239 = lean_unsigned_to_nat(1u); -x_1240 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_1240, 0, x_1238); -lean_ctor_set(x_1240, 1, x_1237); -lean_ctor_set(x_1240, 2, x_1239); -x_1241 = 0; -x_1242 = lean_box(x_1241); -x_1243 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1243, 0, x_3); -lean_ctor_set(x_1243, 1, x_1242); +lean_object* x_61; lean_object* x_62; +x_61 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4; +lean_inc(x_7); +lean_inc(x_6); lean_inc(x_5); -x_1244 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_1240, x_1240, x_1243, x_1238, lean_box(0), lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1240); -if (lean_obj_tag(x_1244) == 0) -{ -lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1351; uint8_t x_1352; -x_1245 = lean_ctor_get(x_1244, 0); -lean_inc(x_1245); -x_1246 = lean_ctor_get(x_1244, 1); -lean_inc(x_1246); -lean_dec(x_1244); -x_1351 = lean_ctor_get(x_1245, 1); -lean_inc(x_1351); -x_1352 = lean_unbox(x_1351); -lean_dec(x_1351); -if (x_1352 == 0) -{ -lean_dec(x_1245); -lean_dec(x_2); -lean_inc(x_1); -x_1247 = x_1; -goto block_1350; +lean_inc(x_4); +lean_inc(x_3); +x_62 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7(x_61, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_63, x_3, x_4, x_5, x_6, x_7, x_64); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_65; } else { -lean_object* x_1353; lean_object* x_1354; -x_1353 = lean_ctor_get(x_1245, 0); -lean_inc(x_1353); -lean_dec(x_1245); -x_1354 = l_Lean_mkAppN(x_2, x_1353); -lean_dec(x_1353); -x_1247 = x_1354; -goto block_1350; -} -block_1350: -{ -lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; uint8_t x_1251; -x_1248 = lean_st_ref_take(x_5, x_1246); -x_1249 = lean_ctor_get(x_1248, 0); -lean_inc(x_1249); -x_1250 = lean_ctor_get(x_1248, 1); -lean_inc(x_1250); -lean_dec(x_1248); -x_1251 = !lean_is_exclusive(x_1249); -if (x_1251 == 0) -{ -lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; size_t x_1255; uint64_t x_1256; uint64_t x_1257; uint64_t x_1258; uint64_t x_1259; uint64_t x_1260; uint64_t x_1261; uint64_t x_1262; uint64_t x_1263; uint64_t x_1264; size_t x_1265; size_t x_1266; size_t x_1267; size_t x_1268; size_t x_1269; lean_object* x_1270; uint8_t x_1271; -x_1252 = lean_ctor_get(x_1249, 0); -x_1253 = lean_ctor_get(x_1249, 1); -x_1254 = lean_array_get_size(x_1253); -x_1255 = lean_ptr_addr(x_1); -x_1256 = lean_usize_to_uint64(x_1255); -x_1257 = 11; -x_1258 = lean_uint64_mix_hash(x_1256, x_1257); -x_1259 = 32; -x_1260 = lean_uint64_shift_right(x_1258, x_1259); -x_1261 = lean_uint64_xor(x_1258, x_1260); -x_1262 = 16; -x_1263 = lean_uint64_shift_right(x_1261, x_1262); -x_1264 = lean_uint64_xor(x_1261, x_1263); -x_1265 = lean_uint64_to_usize(x_1264); -x_1266 = lean_usize_of_nat(x_1254); -lean_dec(x_1254); -x_1267 = 1; -x_1268 = lean_usize_sub(x_1266, x_1267); -x_1269 = lean_usize_land(x_1265, x_1268); -x_1270 = lean_array_uget(x_1253, x_1269); -x_1271 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1270); -if (x_1271 == 0) -{ -lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; uint8_t x_1280; -x_1272 = lean_nat_add(x_1252, x_1239); -lean_dec(x_1252); -lean_inc(x_1247); -x_1273 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1273, 0, x_1); -lean_ctor_set(x_1273, 1, x_1247); -lean_ctor_set(x_1273, 2, x_1270); -x_1274 = lean_array_uset(x_1253, x_1269, x_1273); -x_1275 = lean_unsigned_to_nat(4u); -x_1276 = lean_nat_mul(x_1272, x_1275); -x_1277 = lean_unsigned_to_nat(3u); -x_1278 = lean_nat_div(x_1276, x_1277); -lean_dec(x_1276); -x_1279 = lean_array_get_size(x_1274); -x_1280 = lean_nat_dec_le(x_1278, x_1279); -lean_dec(x_1279); -lean_dec(x_1278); -if (x_1280 == 0) -{ -lean_object* x_1281; lean_object* x_1282; uint8_t x_1283; -x_1281 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1274); -lean_ctor_set(x_1249, 1, x_1281); -lean_ctor_set(x_1249, 0, x_1272); -x_1282 = lean_st_ref_set(x_5, x_1249, x_1250); +uint8_t x_66; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1283 = !lean_is_exclusive(x_1282); -if (x_1283 == 0) +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_62); +if (x_66 == 0) { -lean_object* x_1284; -x_1284 = lean_ctor_get(x_1282, 0); -lean_dec(x_1284); -lean_ctor_set(x_1282, 0, x_1247); -return x_1282; +return x_62; } else { -lean_object* x_1285; lean_object* x_1286; -x_1285 = lean_ctor_get(x_1282, 1); -lean_inc(x_1285); -lean_dec(x_1282); -x_1286 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1286, 0, x_1247); -lean_ctor_set(x_1286, 1, x_1285); -return x_1286; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_62, 0); +x_68 = lean_ctor_get(x_62, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_62); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; } } -else -{ -lean_object* x_1287; uint8_t x_1288; -lean_ctor_set(x_1249, 1, x_1274); -lean_ctor_set(x_1249, 0, x_1272); -x_1287 = lean_st_ref_set(x_5, x_1249, x_1250); -lean_dec(x_5); -x_1288 = !lean_is_exclusive(x_1287); -if (x_1288 == 0) -{ -lean_object* x_1289; -x_1289 = lean_ctor_get(x_1287, 0); -lean_dec(x_1289); -lean_ctor_set(x_1287, 0, x_1247); -return x_1287; } -else -{ -lean_object* x_1290; lean_object* x_1291; -x_1290 = lean_ctor_get(x_1287, 1); -lean_inc(x_1290); -lean_dec(x_1287); -x_1291 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1291, 0, x_1247); -lean_ctor_set(x_1291, 1, x_1290); -return x_1291; } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; uint8_t x_1297; -x_1292 = lean_box(0); -x_1293 = lean_array_uset(x_1253, x_1269, x_1292); -lean_inc(x_1247); -x_1294 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1247, x_1270); -x_1295 = lean_array_uset(x_1293, x_1269, x_1294); -lean_ctor_set(x_1249, 1, x_1295); -x_1296 = lean_st_ref_set(x_5, x_1249, x_1250); -lean_dec(x_5); -x_1297 = !lean_is_exclusive(x_1296); -if (x_1297 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_3, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) { -lean_object* x_1298; -x_1298 = lean_ctor_get(x_1296, 0); -lean_dec(x_1298); -lean_ctor_set(x_1296, 0, x_1247); -return x_1296; -} -else +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_array_get_size(x_13); +x_15 = lean_ptr_addr(x_1); +x_16 = lean_usize_to_uint64(x_15); +x_17 = 11; +x_18 = lean_uint64_mix_hash(x_16, x_17); +x_19 = 32; +x_20 = lean_uint64_shift_right(x_18, x_19); +x_21 = lean_uint64_xor(x_18, x_20); +x_22 = 16; +x_23 = lean_uint64_shift_right(x_21, x_22); +x_24 = lean_uint64_xor(x_21, x_23); +x_25 = lean_uint64_to_usize(x_24); +x_26 = lean_usize_of_nat(x_14); +lean_dec(x_14); +x_27 = 1; +x_28 = lean_usize_sub(x_26, x_27); +x_29 = lean_usize_land(x_25, x_28); +x_30 = lean_array_uget(x_13, x_29); +lean_dec(x_13); +x_31 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(x_1, x_30); +lean_dec(x_30); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_1299; lean_object* x_1300; -x_1299 = lean_ctor_get(x_1296, 1); -lean_inc(x_1299); -lean_dec(x_1296); -x_1300 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1300, 0, x_1247); -lean_ctor_set(x_1300, 1, x_1299); -return x_1300; -} -} +lean_object* x_32; lean_object* x_33; +lean_free_object(x_9); +x_32 = lean_box(0); +x_33 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_32, x_3, x_4, x_5, x_6, x_7, x_12); +return x_33; } else { -lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; size_t x_1304; uint64_t x_1305; uint64_t x_1306; uint64_t x_1307; uint64_t x_1308; uint64_t x_1309; uint64_t x_1310; uint64_t x_1311; uint64_t x_1312; uint64_t x_1313; size_t x_1314; size_t x_1315; size_t x_1316; size_t x_1317; size_t x_1318; lean_object* x_1319; uint8_t x_1320; -x_1301 = lean_ctor_get(x_1249, 0); -x_1302 = lean_ctor_get(x_1249, 1); -lean_inc(x_1302); -lean_inc(x_1301); -lean_dec(x_1249); -x_1303 = lean_array_get_size(x_1302); -x_1304 = lean_ptr_addr(x_1); -x_1305 = lean_usize_to_uint64(x_1304); -x_1306 = 11; -x_1307 = lean_uint64_mix_hash(x_1305, x_1306); -x_1308 = 32; -x_1309 = lean_uint64_shift_right(x_1307, x_1308); -x_1310 = lean_uint64_xor(x_1307, x_1309); -x_1311 = 16; -x_1312 = lean_uint64_shift_right(x_1310, x_1311); -x_1313 = lean_uint64_xor(x_1310, x_1312); -x_1314 = lean_uint64_to_usize(x_1313); -x_1315 = lean_usize_of_nat(x_1303); -lean_dec(x_1303); -x_1316 = 1; -x_1317 = lean_usize_sub(x_1315, x_1316); -x_1318 = lean_usize_land(x_1314, x_1317); -x_1319 = lean_array_uget(x_1302, x_1318); -x_1320 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_1319); -if (x_1320 == 0) -{ -lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; uint8_t x_1329; -x_1321 = lean_nat_add(x_1301, x_1239); -lean_dec(x_1301); -lean_inc(x_1247); -x_1322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1322, 0, x_1); -lean_ctor_set(x_1322, 1, x_1247); -lean_ctor_set(x_1322, 2, x_1319); -x_1323 = lean_array_uset(x_1302, x_1318, x_1322); -x_1324 = lean_unsigned_to_nat(4u); -x_1325 = lean_nat_mul(x_1321, x_1324); -x_1326 = lean_unsigned_to_nat(3u); -x_1327 = lean_nat_div(x_1325, x_1326); -lean_dec(x_1325); -x_1328 = lean_array_get_size(x_1323); -x_1329 = lean_nat_dec_le(x_1327, x_1328); -lean_dec(x_1328); -lean_dec(x_1327); -if (x_1329 == 0) -{ -lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; -x_1330 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_1323); -x_1331 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1331, 0, x_1321); -lean_ctor_set(x_1331, 1, x_1330); -x_1332 = lean_st_ref_set(x_5, x_1331, x_1250); +lean_object* x_34; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_1333 = lean_ctor_get(x_1332, 1); -lean_inc(x_1333); -if (lean_is_exclusive(x_1332)) { - lean_ctor_release(x_1332, 0); - lean_ctor_release(x_1332, 1); - x_1334 = x_1332; -} else { - lean_dec_ref(x_1332); - x_1334 = lean_box(0); -} -if (lean_is_scalar(x_1334)) { - x_1335 = lean_alloc_ctor(0, 2, 0); -} else { - x_1335 = x_1334; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +lean_ctor_set(x_9, 0, x_34); +return x_9; } -lean_ctor_set(x_1335, 0, x_1247); -lean_ctor_set(x_1335, 1, x_1333); -return x_1335; } else { -lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; -x_1336 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1336, 0, x_1321); -lean_ctor_set(x_1336, 1, x_1323); -x_1337 = lean_st_ref_set(x_5, x_1336, x_1250); -lean_dec(x_5); -x_1338 = lean_ctor_get(x_1337, 1); -lean_inc(x_1338); -if (lean_is_exclusive(x_1337)) { - lean_ctor_release(x_1337, 0); - lean_ctor_release(x_1337, 1); - x_1339 = x_1337; -} else { - lean_dec_ref(x_1337); - x_1339 = lean_box(0); -} -if (lean_is_scalar(x_1339)) { - x_1340 = lean_alloc_ctor(0, 2, 0); -} else { - x_1340 = x_1339; -} -lean_ctor_set(x_1340, 0, x_1247); -lean_ctor_set(x_1340, 1, x_1338); -return x_1340; -} -} -else +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; size_t x_39; uint64_t x_40; uint64_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; size_t x_53; lean_object* x_54; lean_object* x_55; +x_35 = lean_ctor_get(x_9, 0); +x_36 = lean_ctor_get(x_9, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_9); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_array_get_size(x_37); +x_39 = lean_ptr_addr(x_1); +x_40 = lean_usize_to_uint64(x_39); +x_41 = 11; +x_42 = lean_uint64_mix_hash(x_40, x_41); +x_43 = 32; +x_44 = lean_uint64_shift_right(x_42, x_43); +x_45 = lean_uint64_xor(x_42, x_44); +x_46 = 16; +x_47 = lean_uint64_shift_right(x_45, x_46); +x_48 = lean_uint64_xor(x_45, x_47); +x_49 = lean_uint64_to_usize(x_48); +x_50 = lean_usize_of_nat(x_38); +lean_dec(x_38); +x_51 = 1; +x_52 = lean_usize_sub(x_50, x_51); +x_53 = lean_usize_land(x_49, x_52); +x_54 = lean_array_uget(x_37, x_53); +lean_dec(x_37); +x_55 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(x_1, x_54); +lean_dec(x_54); +if (lean_obj_tag(x_55) == 0) { -lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; -x_1341 = lean_box(0); -x_1342 = lean_array_uset(x_1302, x_1318, x_1341); -lean_inc(x_1247); -x_1343 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_1247, x_1319); -x_1344 = lean_array_uset(x_1342, x_1318, x_1343); -x_1345 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1345, 0, x_1301); -lean_ctor_set(x_1345, 1, x_1344); -x_1346 = lean_st_ref_set(x_5, x_1345, x_1250); -lean_dec(x_5); -x_1347 = lean_ctor_get(x_1346, 1); -lean_inc(x_1347); -if (lean_is_exclusive(x_1346)) { - lean_ctor_release(x_1346, 0); - lean_ctor_release(x_1346, 1); - x_1348 = x_1346; -} else { - lean_dec_ref(x_1346); - x_1348 = lean_box(0); -} -if (lean_is_scalar(x_1348)) { - x_1349 = lean_alloc_ctor(0, 2, 0); -} else { - x_1349 = x_1348; -} -lean_ctor_set(x_1349, 0, x_1247); -lean_ctor_set(x_1349, 1, x_1347); -return x_1349; -} -} -} +lean_object* x_56; lean_object* x_57; +x_56 = lean_box(0); +x_57 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_56, x_3, x_4, x_5, x_6, x_7, x_36); +return x_57; } else { -uint8_t x_1355; +lean_object* x_58; lean_object* x_59; +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_1355 = !lean_is_exclusive(x_1244); -if (x_1355 == 0) -{ -return x_1244; -} -else -{ -lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; -x_1356 = lean_ctor_get(x_1244, 0); -x_1357 = lean_ctor_get(x_1244, 1); -lean_inc(x_1357); -lean_inc(x_1356); -lean_dec(x_1244); -x_1358 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1358, 0, x_1356); -lean_ctor_set(x_1358, 1, x_1357); -return x_1358; -} -} +x_58 = lean_ctor_get(x_55, 0); +lean_inc(x_58); +lean_dec(x_55); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_36); +return x_59; } } } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -if (lean_obj_tag(x_2) == 0) +uint8_t x_9; +x_9 = l_Lean_Expr_isApp(x_1); +if (x_9 == 0) { -lean_object* x_3; -x_3 = lean_box(0); -return x_3; -} -else +uint8_t x_10; +x_10 = l_Lean_Expr_isForall(x_1); +if (x_10 == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint8_t x_9; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_2, 2); -x_7 = lean_ptr_addr(x_4); -x_8 = lean_ptr_addr(x_1); -x_9 = lean_usize_dec_eq(x_7, x_8); -if (x_9 == 0) +uint8_t x_11; +x_11 = l_Lean_Expr_isProj(x_1); +if (x_11 == 0) { -x_2 = x_6; -goto _start; +lean_object* x_12; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_8); +return x_12; } else { -lean_object* x_11; -lean_inc(x_5); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_5); -return x_11; -} -} +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(0); +x_14 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8); +return x_14; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_8); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_9 = lean_unsigned_to_nat(0u); -x_10 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_9); -x_11 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1; -lean_inc(x_10); -x_12 = lean_mk_array(x_10, x_11); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_10, x_13); -lean_dec(x_10); -lean_inc(x_1); -x_15 = l_Lean_Expr_withAppAux___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__9(x_1, x_1, x_12, x_14, x_3, x_4, x_5, x_6, x_7, x_8); -return x_15; +lean_object* x_17; lean_object* x_18; +x_17 = lean_box(0); +x_18 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(x_1, x_17, x_3, x_4, x_5, x_6, x_7, x_8); +return x_18; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1() { +} +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1() { _start: { lean_object* x_1; @@ -4887,7 +1928,7 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2() { _start: { lean_object* x_1; @@ -4895,7 +1936,7 @@ x_1 = lean_mk_string_unchecked("Grind", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3() { +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3() { _start: { lean_object* x_1; @@ -4903,28 +1944,28 @@ x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1; -x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2; -x_3 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3; +x_1 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1; +x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2; +x_3 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; +x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; @@ -4938,7 +1979,7 @@ lean_inc(x_10); x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); -x_12 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5; +x_12 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5; lean_inc(x_1); x_13 = l_Lean_mkAppB(x_12, x_10, x_1); x_14 = lean_st_ref_take(x_3, x_11); @@ -4971,7 +2012,7 @@ x_33 = 1; x_34 = lean_usize_sub(x_32, x_33); x_35 = lean_usize_land(x_31, x_34); x_36 = lean_array_uget(x_19, x_35); -x_37 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_36); +x_37 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_36); if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; @@ -4996,7 +2037,7 @@ lean_dec(x_45); if (x_47 == 0) { lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_41); +x_48 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_41); lean_ctor_set(x_15, 1, x_48); lean_ctor_set(x_15, 0, x_39); x_49 = lean_st_ref_set(x_3, x_15, x_16); @@ -5055,7 +2096,7 @@ lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean x_59 = lean_box(0); x_60 = lean_array_uset(x_19, x_35, x_59); lean_inc(x_13); -x_61 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_13, x_36); +x_61 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_13, x_36); x_62 = lean_array_uset(x_60, x_35, x_61); lean_ctor_set(x_15, 1, x_62); x_63 = lean_st_ref_set(x_3, x_15, x_16); @@ -5107,7 +2148,7 @@ x_83 = 1; x_84 = lean_usize_sub(x_82, x_83); x_85 = lean_usize_land(x_81, x_84); x_86 = lean_array_uget(x_69, x_85); -x_87 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_86); +x_87 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_86); if (x_87 == 0) { lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; @@ -5132,7 +2173,7 @@ lean_dec(x_95); if (x_97 == 0) { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_98 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__4(x_91); +x_98 = l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_91); x_99 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_99, 0, x_89); lean_ctor_set(x_99, 1, x_98); @@ -5189,7 +2230,7 @@ lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; x_109 = lean_box(0); x_110 = lean_array_uset(x_69, x_85, x_109); lean_inc(x_13); -x_111 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_13, x_86); +x_111 = l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__6(x_1, x_13, x_86); x_112 = lean_array_uset(x_110, x_85, x_111); x_113 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_113, 0, x_68); @@ -5241,7 +2282,7 @@ return x_121; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; uint8_t x_10; @@ -5281,7 +2322,7 @@ if (lean_obj_tag(x_31) == 0) lean_object* x_32; lean_object* x_33; lean_free_object(x_9); x_32 = lean_box(0); -x_33 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_1, x_32, x_3, x_4, x_5, x_6, x_7, x_12); +x_33 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(x_1, x_32, x_3, x_4, x_5, x_6, x_7, x_12); return x_33; } else @@ -5335,7 +2376,7 @@ if (lean_obj_tag(x_55) == 0) { lean_object* x_56; lean_object* x_57; x_56 = lean_box(0); -x_57 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_1, x_56, x_3, x_4, x_5, x_6, x_7, x_36); +x_57 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(x_1, x_56, x_3, x_4, x_5, x_6, x_7, x_36); return x_57; } else @@ -5357,43 +2398,6 @@ return x_59; } } } -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.MarkNestedProofs", 39, 39); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Meta.Grind.markNestedProofsImpl.visit", 42, 42); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1; -x_2 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2; -x_3 = lean_unsigned_to_nat(28u); -x_4 = lean_unsigned_to_nat(20u); -x_5 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { @@ -5413,393 +2417,34 @@ x_10 = lean_unbox(x_9); lean_dec(x_9); if (x_10 == 0) { -switch (lean_obj_tag(x_1)) { -case 0: -{ lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_1); x_11 = lean_ctor_get(x_8, 1); lean_inc(x_11); lean_dec(x_8); -x_12 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4; -x_13 = l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_12, x_2, x_3, x_4, x_5, x_6, x_11); +x_12 = lean_box(0); +x_13 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_11); return x_13; } -case 5: -{ -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = lean_ctor_get(x_8, 1); -lean_inc(x_14); -lean_dec(x_8); -x_15 = lean_st_ref_get(x_2, x_14); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; uint64_t x_22; uint64_t x_23; uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; uint64_t x_28; uint64_t x_29; uint64_t x_30; size_t x_31; size_t x_32; size_t x_33; size_t x_34; size_t x_35; lean_object* x_36; lean_object* x_37; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_array_get_size(x_19); -x_21 = lean_ptr_addr(x_1); -x_22 = lean_usize_to_uint64(x_21); -x_23 = 11; -x_24 = lean_uint64_mix_hash(x_22, x_23); -x_25 = 32; -x_26 = lean_uint64_shift_right(x_24, x_25); -x_27 = lean_uint64_xor(x_24, x_26); -x_28 = 16; -x_29 = lean_uint64_shift_right(x_27, x_28); -x_30 = lean_uint64_xor(x_27, x_29); -x_31 = lean_uint64_to_usize(x_30); -x_32 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_33 = 1; -x_34 = lean_usize_sub(x_32, x_33); -x_35 = lean_usize_land(x_31, x_34); -x_36 = lean_array_uget(x_19, x_35); -lean_dec(x_19); -x_37 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(x_1, x_36); -lean_dec(x_36); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; -lean_free_object(x_15); -x_38 = lean_box(0); -x_39 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_38, x_2, x_3, x_4, x_5, x_6, x_18); -return x_39; -} -else -{ -lean_object* x_40; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_40 = lean_ctor_get(x_37, 0); -lean_inc(x_40); -lean_dec(x_37); -lean_ctor_set(x_15, 0, x_40); -return x_15; -} -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; size_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; uint64_t x_50; uint64_t x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; size_t x_55; size_t x_56; size_t x_57; size_t x_58; size_t x_59; lean_object* x_60; lean_object* x_61; -x_41 = lean_ctor_get(x_15, 0); -x_42 = lean_ctor_get(x_15, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_15); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = lean_array_get_size(x_43); -x_45 = lean_ptr_addr(x_1); -x_46 = lean_usize_to_uint64(x_45); -x_47 = 11; -x_48 = lean_uint64_mix_hash(x_46, x_47); -x_49 = 32; -x_50 = lean_uint64_shift_right(x_48, x_49); -x_51 = lean_uint64_xor(x_48, x_50); -x_52 = 16; -x_53 = lean_uint64_shift_right(x_51, x_52); -x_54 = lean_uint64_xor(x_51, x_53); -x_55 = lean_uint64_to_usize(x_54); -x_56 = lean_usize_of_nat(x_44); -lean_dec(x_44); -x_57 = 1; -x_58 = lean_usize_sub(x_56, x_57); -x_59 = lean_usize_land(x_55, x_58); -x_60 = lean_array_uget(x_43, x_59); -lean_dec(x_43); -x_61 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10(x_1, x_60); -lean_dec(x_60); -if (lean_obj_tag(x_61) == 0) -{ -lean_object* x_62; lean_object* x_63; -x_62 = lean_box(0); -x_63 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_62, x_2, x_3, x_4, x_5, x_6, x_42); -return x_63; -} -else -{ -lean_object* x_64; lean_object* x_65; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_64 = lean_ctor_get(x_61, 0); -lean_inc(x_64); -lean_dec(x_61); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_42); -return x_65; -} -} -} -case 10: -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_8, 1); -lean_inc(x_66); -lean_dec(x_8); -x_67 = lean_ctor_get(x_1, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_1, 1); -lean_inc(x_68); -lean_inc(x_68); -x_69 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_68, x_2, x_3, x_4, x_5, x_6, x_66); -if (lean_obj_tag(x_69) == 0) -{ -uint8_t x_70; -x_70 = !lean_is_exclusive(x_69); -if (x_70 == 0) -{ -lean_object* x_71; size_t x_72; size_t x_73; uint8_t x_74; -x_71 = lean_ctor_get(x_69, 0); -x_72 = lean_ptr_addr(x_68); -lean_dec(x_68); -x_73 = lean_ptr_addr(x_71); -x_74 = lean_usize_dec_eq(x_72, x_73); -if (x_74 == 0) -{ -lean_object* x_75; -lean_dec(x_1); -x_75 = l_Lean_Expr_mdata___override(x_67, x_71); -lean_ctor_set(x_69, 0, x_75); -return x_69; -} -else -{ -lean_dec(x_71); -lean_dec(x_67); -lean_ctor_set(x_69, 0, x_1); -return x_69; -} -} -else -{ -lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; uint8_t x_80; -x_76 = lean_ctor_get(x_69, 0); -x_77 = lean_ctor_get(x_69, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_69); -x_78 = lean_ptr_addr(x_68); -lean_dec(x_68); -x_79 = lean_ptr_addr(x_76); -x_80 = lean_usize_dec_eq(x_78, x_79); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -lean_dec(x_1); -x_81 = l_Lean_Expr_mdata___override(x_67, x_76); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_77); -return x_82; -} -else -{ -lean_object* x_83; -lean_dec(x_76); -lean_dec(x_67); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_1); -lean_ctor_set(x_83, 1, x_77); -return x_83; -} -} -} -else -{ -uint8_t x_84; -lean_dec(x_68); -lean_dec(x_67); -lean_dec(x_1); -x_84 = !lean_is_exclusive(x_69); -if (x_84 == 0) -{ -return x_69; -} -else -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_69, 0); -x_86 = lean_ctor_get(x_69, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_69); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; -} -} -} -case 11: -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_88 = lean_ctor_get(x_8, 1); -lean_inc(x_88); -lean_dec(x_8); -x_89 = lean_ctor_get(x_1, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_1, 1); -lean_inc(x_90); -x_91 = lean_ctor_get(x_1, 2); -lean_inc(x_91); -lean_inc(x_91); -x_92 = l_Lean_Meta_Grind_markNestedProofsImpl_visit(x_91, x_2, x_3, x_4, x_5, x_6, x_88); -if (lean_obj_tag(x_92) == 0) -{ -uint8_t x_93; -x_93 = !lean_is_exclusive(x_92); -if (x_93 == 0) -{ -lean_object* x_94; size_t x_95; size_t x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_92, 0); -x_95 = lean_ptr_addr(x_91); -lean_dec(x_91); -x_96 = lean_ptr_addr(x_94); -x_97 = lean_usize_dec_eq(x_95, x_96); -if (x_97 == 0) -{ -lean_object* x_98; -lean_dec(x_1); -x_98 = l_Lean_Expr_proj___override(x_89, x_90, x_94); -lean_ctor_set(x_92, 0, x_98); -return x_92; -} -else -{ -lean_dec(x_94); -lean_dec(x_90); -lean_dec(x_89); -lean_ctor_set(x_92, 0, x_1); -return x_92; -} -} -else -{ -lean_object* x_99; lean_object* x_100; size_t x_101; size_t x_102; uint8_t x_103; -x_99 = lean_ctor_get(x_92, 0); -x_100 = lean_ctor_get(x_92, 1); -lean_inc(x_100); -lean_inc(x_99); -lean_dec(x_92); -x_101 = lean_ptr_addr(x_91); -lean_dec(x_91); -x_102 = lean_ptr_addr(x_99); -x_103 = lean_usize_dec_eq(x_101, x_102); -if (x_103 == 0) -{ -lean_object* x_104; lean_object* x_105; -lean_dec(x_1); -x_104 = l_Lean_Expr_proj___override(x_89, x_90, x_99); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_100); -return x_105; -} -else -{ -lean_object* x_106; -lean_dec(x_99); -lean_dec(x_90); -lean_dec(x_89); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_1); -lean_ctor_set(x_106, 1, x_100); -return x_106; -} -} -} -else -{ -uint8_t x_107; -lean_dec(x_91); -lean_dec(x_90); -lean_dec(x_89); -lean_dec(x_1); -x_107 = !lean_is_exclusive(x_92); -if (x_107 == 0) -{ -return x_92; -} -else -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_92, 0); -x_109 = lean_ctor_get(x_92, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_92); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; -} -} -} -default: -{ -uint8_t x_111; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_111 = !lean_is_exclusive(x_8); -if (x_111 == 0) -{ -lean_object* x_112; -x_112 = lean_ctor_get(x_8, 0); -lean_dec(x_112); -lean_ctor_set(x_8, 0, x_1); -return x_8; -} else { -lean_object* x_113; lean_object* x_114; -x_113 = lean_ctor_get(x_8, 1); -lean_inc(x_113); -lean_dec(x_8); -x_114 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_114, 0, x_1); -lean_ctor_set(x_114, 1, x_113); -return x_114; -} -} -} -} -else -{ -uint8_t x_115; -x_115 = !lean_is_exclusive(x_8); -if (x_115 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_8); +if (x_14 == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; -x_116 = lean_ctor_get(x_8, 1); -x_117 = lean_ctor_get(x_8, 0); -lean_dec(x_117); -x_118 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; -x_119 = l_Lean_Expr_isAppOf(x_1, x_118); -if (x_119 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_8, 1); +x_16 = lean_ctor_get(x_8, 0); +lean_dec(x_16); +x_17 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4; +x_18 = l_Lean_Expr_isAppOf(x_1, x_17); +if (x_18 == 0) { -lean_object* x_120; lean_object* x_121; +lean_object* x_19; lean_object* x_20; lean_free_object(x_8); -x_120 = lean_box(0); -x_121 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_120, x_2, x_3, x_4, x_5, x_6, x_116); +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(x_1, x_19, x_2, x_3, x_4, x_5, x_6, x_15); lean_dec(x_2); -return x_121; +return x_20; } else { @@ -5814,85 +2459,85 @@ return x_8; } else { -lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_122 = lean_ctor_get(x_8, 1); -lean_inc(x_122); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_8, 1); +lean_inc(x_21); lean_dec(x_8); -x_123 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4; -x_124 = l_Lean_Expr_isAppOf(x_1, x_123); -if (x_124 == 0) +x_22 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4; +x_23 = l_Lean_Expr_isAppOf(x_1, x_22); +if (x_23 == 0) { -lean_object* x_125; lean_object* x_126; -x_125 = lean_box(0); -x_126 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_125, x_2, x_3, x_4, x_5, x_6, x_122); +lean_object* x_24; lean_object* x_25; +x_24 = lean_box(0); +x_25 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(x_1, x_24, x_2, x_3, x_4, x_5, x_6, x_21); lean_dec(x_2); -return x_126; +return x_25; } else { -lean_object* x_127; +lean_object* x_26; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_127 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_127, 0, x_1); -lean_ctor_set(x_127, 1, x_122); -return x_127; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_21); +return x_26; } } } } else { -uint8_t x_128; +uint8_t x_27; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_128 = !lean_is_exclusive(x_8); -if (x_128 == 0) +x_27 = !lean_is_exclusive(x_8); +if (x_27 == 0) { return x_8; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_129 = lean_ctor_get(x_8, 0); -x_130 = lean_ctor_get(x_8, 1); -lean_inc(x_130); -lean_inc(x_129); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_8, 0); +x_29 = lean_ctor_get(x_8, 1); +lean_inc(x_29); +lean_inc(x_28); lean_dec(x_8); -x_131 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_131, 0, x_129); -lean_ctor_set(x_131, 1, x_130); -return x_131; +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_13; -x_13 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); -return x_13; +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__3(x_1, x_2); +lean_object* x_13; +x_13 = l_Std_Range_forIn_x27_loop___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_2); lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +return x_13; } } LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__10___boxed(lean_object* x_1, lean_object* x_2) { @@ -5910,25 +2555,66 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__ { lean_object* x_9; x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_4); +lean_dec(x_4); +x_15 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_1, x_2, x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); lean_dec(x_2); return x_9; @@ -6049,30 +2735,30 @@ lean_dec_ref(res); res = initialize_Lean_Meta_InferType(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1 = _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__1); -l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2 = _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2(); -lean_mark_persistent(l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__1___closed__2); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__1___closed__1); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__1); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__2); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__3); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__4); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__2___closed__5); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__1); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__2); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__3); -l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___closed__4); +l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1 = _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__1); +l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2 = _init_l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Meta_Grind_markNestedProofsImpl_visit___spec__7___closed__2); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__1); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__2); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__3); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__4); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__3___closed__5); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__1); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__2); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__3); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__4); +l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5 = _init_l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl_visit___lambda__6___closed__5); l_Lean_Meta_Grind_markNestedProofsImpl___closed__1 = _init_l_Lean_Meta_Grind_markNestedProofsImpl___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_markNestedProofsImpl___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c index a7040611d752..5cfa35c8a3b9 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proj.c @@ -37,7 +37,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq(lean_object*, lean_ob lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__3; static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4; -lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_Grind_propagateProjEq___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -45,19 +44,21 @@ lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, lean_object*, lean_ob uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; lean_object* l_Lean_MessageData_ofExpr(lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -270,7 +271,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateProjEq___lambda__3(lean_obje { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__4; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_addCongrTable___spec__8(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_unbox(x_15); @@ -291,53 +292,126 @@ uint8_t x_20; x_20 = !lean_is_exclusive(x_14); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_14, 1); x_22 = lean_ctor_get(x_14, 0); lean_dec(x_22); +x_23 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); lean_inc(x_3); -x_23 = l_Lean_MessageData_ofExpr(x_3); -x_24 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; +x_25 = l_Lean_MessageData_ofExpr(x_3); +x_26 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; lean_ctor_set_tag(x_14, 7); -lean_ctor_set(x_14, 1, x_23); -lean_ctor_set(x_14, 0, x_24); -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); -lean_dec(x_27); -return x_29; +lean_ctor_set(x_14, 1, x_25); +lean_ctor_set(x_14, 0, x_26); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_14); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_29); +return x_31; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); +uint8_t x_32; +lean_free_object(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) +{ +return x_23; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_23, 0); +x_34 = lean_ctor_get(x_23, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_23); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); lean_dec(x_14); +x_37 = l_Lean_Meta_Grind_updateLastTag(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_36); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); lean_inc(x_3); -x_31 = l_Lean_MessageData_ofExpr(x_3); -x_32 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_addTrace___at_Lean_Meta_Grind_addCongrTable___spec__9(x_13, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); -lean_dec(x_36); -return x_38; +x_39 = l_Lean_MessageData_ofExpr(x_3); +x_40 = l_Lean_Meta_Grind_propagateProjEq___lambda__3___closed__6; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +x_43 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_13, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_propagateProjEq___lambda__2(x_1, x_2, x_3, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); +lean_dec(x_44); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +x_47 = lean_ctor_get(x_37, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_37, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_49 = x_37; +} else { + lean_dec_ref(x_37); + x_49 = lean_box(0); +} +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(1, 2, 0); +} else { + x_50 = x_49; +} +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c index 3378f014e3c2..8a674a81ce00 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Proof.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Proof -// Imports: Lean.Meta.Sorry Lean.Meta.Tactic.Grind.Types +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Grind.Types #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,6 +15,7 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__3; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__4; @@ -24,51 +25,60 @@ lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___lambda__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__5; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__2; +lean_object* l_Lean_mkApp7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkRefl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4; uint8_t l_Lean_Expr_isApp(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__4; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6; static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__4; lean_object* l_Lean_Meta_mkHEqOfEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5___closed__2; +uint8_t l_Lean_Meta_Grind_hasSameRoot(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1; lean_object* l_Lean_Expr_appArg_x21(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_grind_mk_heq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans_x27(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -77,13 +87,18 @@ LEAN_EXPORT lean_object* lean_grind_mk_eq_proof(lean_object*, lean_object*, lean LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__1; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getFunInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqNDRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -91,14 +106,19 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_G lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__1; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__6___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkCongrArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11; extern lean_object* l_Lean_levelZero; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_realizeEqProof(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7; extern lean_object* l_Lean_instInhabitedExpr; @@ -107,16 +127,18 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_flipProof(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); lean_object* l_Lean_Meta_Grind_mkHCongrWithArity(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__3; uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__6(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_mkEqProofImpl___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,20 +146,23 @@ static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_ lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5; lean_object* l_Lean_Meta_mkCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_constLevels_x21(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern uint8_t l_Lean_Meta_instInhabitedCongrArgKind; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_realizeEqProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); @@ -145,21 +170,26 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___spec__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2; lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1(lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1; lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__4; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); @@ -171,12 +201,16 @@ lean_object* l_Lean_Meta_mkEqOfHEq(lean_object*, lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkTrans(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1(lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9___closed__1; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkRefl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEqTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -187,15 +221,14 @@ lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7; static lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__2; lean_object* l_Lean_Meta_mkHEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); @@ -5453,7 +5486,7 @@ lean_dec(x_1); return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; lean_object* x_15; uint8_t x_16; @@ -5492,7 +5525,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_21 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(x_1, x_19, x_20, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_21 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(x_1, x_19, x_20, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_20); lean_dec(x_19); if (lean_obj_tag(x_21) == 0) @@ -5693,7 +5726,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; -x_3 = lean_unsigned_to_nat(212u); +x_3 = lean_unsigned_to_nat(234u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -5706,7 +5739,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__4; -x_3 = lean_unsigned_to_nat(215u); +x_3 = lean_unsigned_to_nat(237u); x_4 = lean_unsigned_to_nat(72u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -6224,7 +6257,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(193u); +x_3 = lean_unsigned_to_nat(215u); x_4 = lean_unsigned_to_nat(35u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -6237,7 +6270,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(194u); +x_3 = lean_unsigned_to_nat(216u); x_4 = lean_unsigned_to_nat(29u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7023,167 +7056,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = lean_apply_10(x_1, x_6, x_2, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___lambda__1), 11, 5); -lean_closure_set(x_15, 0, x_4); -lean_closure_set(x_15, 1, x_6); -lean_closure_set(x_15, 2, x_7); -lean_closure_set(x_15, 3, x_8); -lean_closure_set(x_15, 4, x_9); -x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_15, x_5, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_16) == 0) -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -return x_16; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_16); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; -} -} -else -{ -uint8_t x_21; -x_21 = !lean_is_exclusive(x_16); -if (x_21 == 0) -{ -return x_16; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_16, 0); -x_23 = lean_ctor_get(x_16, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_16); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___boxed), 14, 0); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_13); -x_15 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1; -lean_inc(x_14); -x_16 = lean_mk_array(x_14, x_15); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_14, x_17); -lean_dec(x_14); -x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_1, x_16, x_18); -lean_inc(x_3); -x_20 = l_Lean_mkAppN(x_3, x_19); -lean_dec(x_19); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_21 = l_Lean_Meta_mkHEq(x_2, x_20, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_3); -lean_ctor_set(x_25, 1, x_24); -x_26 = lean_array_mk(x_25); -x_27 = 0; -x_28 = 1; -x_29 = 1; -x_30 = l_Lean_Meta_mkLambdaFVars(x_26, x_22, x_27, x_28, x_27, x_29, x_8, x_9, x_10, x_11, x_23); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_26); -return x_30; -} -else -{ -uint8_t x_31; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -x_31 = !lean_is_exclusive(x_21); -if (x_31 == 0) -{ -return x_21; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_21, 0); -x_33 = lean_ctor_get(x_21, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_21); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; -} -} -} -} static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("thm.argKinds.size == numArgs\n ", 35, 35); +x_1 = lean_mk_string_unchecked("rhs.getAppNumArgs == numArgs\n ", 33, 33); return x_1; } } @@ -7211,8 +7088,8 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__3; -x_3 = lean_unsigned_to_nat(143u); -x_4 = lean_unsigned_to_nat(6u); +x_3 = lean_unsigned_to_nat(182u); +x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -7222,60 +7099,11 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("x", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("rhs.getAppNumArgs == numArgs\n ", 33, 33); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__3; -x_3 = lean_unsigned_to_nat(136u); -x_4 = lean_unsigned_to_nat(4u); -x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6() { _start: { lean_object* x_1; @@ -7283,7 +7111,7 @@ x_1 = lean_mk_string_unchecked("Grind", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7() { _start: { lean_object* x_1; @@ -7291,13 +7119,13 @@ x_1 = lean_mk_string_unchecked("nestedProof", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11; -x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; +x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } @@ -7305,293 +7133,197 @@ return x_4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_87; uint8_t x_88; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_58; uint8_t x_59; x_13 = l_Lean_Expr_getAppFn(x_1); x_14 = l_Lean_Expr_getAppFn(x_2); x_15 = lean_unsigned_to_nat(0u); x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); -x_87 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_15); -x_88 = lean_nat_dec_eq(x_87, x_16); -lean_dec(x_87); -if (x_88 == 0) +x_58 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_15); +x_59 = lean_nat_dec_eq(x_58, x_16); +lean_dec(x_58); +if (x_59 == 0) { -lean_object* x_89; lean_object* x_90; +lean_object* x_60; lean_object* x_61; lean_dec(x_16); lean_dec(x_14); lean_dec(x_13); lean_dec(x_2); lean_dec(x_1); -x_89 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9; -x_90 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_89, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_90; +x_60 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__4; +x_61 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_60, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_61; } else { -lean_object* x_91; uint8_t x_92; -x_91 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13; -x_92 = l_Lean_Expr_isConstOf(x_13, x_91); -if (x_92 == 0) +lean_object* x_62; uint8_t x_63; +x_62 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8; +x_63 = l_Lean_Expr_isConstOf(x_13, x_62); +if (x_63 == 0) { -lean_object* x_93; -x_93 = lean_box(0); -x_17 = x_93; -goto block_86; +lean_object* x_64; +x_64 = lean_box(0); +x_17 = x_64; +goto block_57; } else { -uint8_t x_94; -x_94 = l_Lean_Expr_isConstOf(x_14, x_91); -if (x_94 == 0) +uint8_t x_65; +x_65 = l_Lean_Expr_isConstOf(x_14, x_62); +if (x_65 == 0) { -lean_object* x_95; -x_95 = lean_box(0); -x_17 = x_95; -goto block_86; +lean_object* x_66; +x_66 = lean_box(0); +x_17 = x_66; +goto block_57; } else { -lean_object* x_96; uint8_t x_97; -x_96 = lean_unsigned_to_nat(2u); -x_97 = lean_nat_dec_eq(x_16, x_96); -if (x_97 == 0) +lean_object* x_67; uint8_t x_68; +x_67 = lean_unsigned_to_nat(2u); +x_68 = lean_nat_dec_eq(x_16, x_67); +if (x_68 == 0) { -lean_object* x_98; -x_98 = lean_box(0); -x_17 = x_98; -goto block_86; +lean_object* x_69; +x_69 = lean_box(0); +x_17 = x_69; +goto block_57; } else { -lean_object* x_99; +lean_object* x_70; lean_dec(x_16); lean_dec(x_14); lean_dec(x_13); -x_99 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_70 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_2); lean_dec(x_1); -return x_99; +return x_70; } } } } -block_86: +block_57: { -lean_object* x_18; +lean_object* x_18; uint8_t x_19; lean_dec(x_17); +x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; +x_19 = l_Lean_Expr_isConstOf(x_13, x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_13); -x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(x_1, x_2, x_13, x_14, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_unbox(x_19); -lean_dec(x_19); -if (x_20 == 0) +x_20 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(x_1, x_2, x_13, x_14, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_16); +lean_dec(x_14); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 1); +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); -lean_dec(x_18); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_16); -lean_inc(x_13); -x_22 = l_Lean_Meta_Grind_mkHCongrWithArity(x_13, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -if (lean_obj_tag(x_22) == 0) +x_22 = lean_unbox(x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_23 = lean_ctor_get(x_22, 0); +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_23, 2); +lean_dec(x_20); +x_24 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_20, 1); lean_inc(x_25); -x_26 = lean_array_get_size(x_25); -lean_dec(x_25); -x_27 = lean_nat_dec_eq(x_26, x_16); -lean_dec(x_26); -if (x_27 == 0) +lean_dec(x_20); +x_26 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +lean_dec(x_2); +lean_dec(x_1); +return x_26; +} +} +else { -lean_object* x_28; lean_object* x_29; -lean_dec(x_23); -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_27; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_28 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__4; -x_29 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); -return x_29; +x_27 = !lean_is_exclusive(x_20); +if (x_27 == 0) +{ +return x_20; } else { -lean_object* x_30; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_30 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(x_23, x_1, x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_24); -lean_dec(x_16); -lean_dec(x_23); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_13, x_14); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_34 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; -x_35 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_34, x_10, x_11, x_32); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_13); -x_38 = lean_infer_type(x_13, x_8, x_9, x_10, x_11, x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; uint8_t x_43; lean_object* x_44; -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___boxed), 12, 2); -lean_closure_set(x_41, 0, x_2); -lean_closure_set(x_41, 1, x_1); -x_42 = 0; -x_43 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_44 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg(x_36, x_42, x_39, x_41, x_43, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); -if (lean_obj_tag(x_44) == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_20, 0); +x_29 = lean_ctor_get(x_20, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_20); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +else { -lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = 0; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_48 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_13, x_14, x_47, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_46); -if (lean_obj_tag(x_48) == 0) +uint8_t x_31; +x_31 = l_Lean_Expr_isConstOf(x_14, x_18); +if (x_31 == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); +lean_object* x_32; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_51 = l_Lean_Meta_mkEqNDRec(x_45, x_31, x_49, x_8, x_9, x_10, x_11, x_50); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(x_52, x_3, x_8, x_9, x_10, x_11, x_53); -return x_54; -} -else -{ -uint8_t x_55; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_55 = !lean_is_exclusive(x_51); -if (x_55 == 0) -{ -return x_51; -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_51, 0); -x_57 = lean_ctor_get(x_51, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_51); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -return x_58; -} -} -} -else +x_32 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(x_1, x_2, x_13, x_14, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_16); +lean_dec(x_14); +if (lean_obj_tag(x_32) == 0) { -uint8_t x_59; -lean_dec(x_45); -lean_dec(x_31); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_59 = !lean_is_exclusive(x_48); -if (x_59 == 0) +lean_object* x_33; uint8_t x_34; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_unbox(x_33); +lean_dec(x_33); +if (x_34 == 0) { -return x_48; +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_dec(x_32); +x_36 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_35); +return x_36; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_48, 0); -x_61 = lean_ctor_get(x_48, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_48); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; -} +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_32, 1); +lean_inc(x_37); +lean_dec(x_32); +x_38 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_37); +lean_dec(x_2); +lean_dec(x_1); +return x_38; } } else { -uint8_t x_63; -lean_dec(x_31); -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_39; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7600,83 +7332,74 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_63 = !lean_is_exclusive(x_44); -if (x_63 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_39 = !lean_is_exclusive(x_32); +if (x_39 == 0) { -return x_44; +return x_32; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_44, 0); -x_65 = lean_ctor_get(x_44, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_44); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_32, 0); +x_41 = lean_ctor_get(x_32, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_32); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } else { -uint8_t x_67; -lean_dec(x_36); -lean_dec(x_31); +lean_object* x_43; uint8_t x_44; +x_43 = lean_unsigned_to_nat(3u); +x_44 = lean_nat_dec_eq(x_16, x_43); +if (x_44 == 0) +{ +lean_object* x_45; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_45 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isCongrDefaultProofTarget(x_1, x_2, x_13, x_14, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_16); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_67 = !lean_is_exclusive(x_38); -if (x_67 == 0) +if (lean_obj_tag(x_45) == 0) { -return x_38; -} -else +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_38, 0); -x_69 = lean_ctor_get(x_38, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_38); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; -} -} +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +x_49 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_48); +return x_49; } else { -lean_object* x_71; -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); +lean_dec(x_45); +x_51 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_50); lean_dec(x_2); lean_dec(x_1); -x_71 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(x_31, x_3, x_8, x_9, x_10, x_11, x_32); -return x_71; +return x_51; } } else { -uint8_t x_72; -lean_dec(x_14); -lean_dec(x_13); +uint8_t x_52; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7687,331 +7410,329 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_72 = !lean_is_exclusive(x_30); -if (x_72 == 0) +x_52 = !lean_is_exclusive(x_45); +if (x_52 == 0) { -return x_30; +return x_45; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_30, 0); -x_74 = lean_ctor_get(x_30, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_30); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -return x_75; -} +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_45, 0); +x_54 = lean_ctor_get(x_45, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_45); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } } else { -uint8_t x_76; +lean_object* x_56; lean_dec(x_16); lean_dec(x_14); lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_76 = !lean_is_exclusive(x_22); -if (x_76 == 0) -{ -return x_22; +x_56 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_56; } -else -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_22, 0); -x_78 = lean_ctor_get(x_22, 1); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_22); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; } } } -else +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_80; lean_object* x_81; -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_13); -x_80 = lean_ctor_get(x_18, 1); -lean_inc(x_80); -lean_dec(x_18); -x_81 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_80); -lean_dec(x_2); -lean_dec(x_1); -return x_81; +lean_object* x_12; +x_12 = lean_apply_10(x_1, x_6, x_2, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11); +return x_12; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -uint8_t x_82; -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_82 = !lean_is_exclusive(x_18); -if (x_82 == 0) +lean_object* x_15; lean_object* x_16; +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___lambda__1), 11, 5); +lean_closure_set(x_15, 0, x_4); +lean_closure_set(x_15, 1, x_6); +lean_closure_set(x_15, 2, x_7); +lean_closure_set(x_15, 3, x_8); +lean_closure_set(x_15, 4, x_9); +x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_15, x_5, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) { -return x_18; +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +return x_16; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_18, 0); -x_84 = lean_ctor_get(x_18, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_18); -x_85 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -return x_85; -} -} +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) +{ +return x_16; } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); -return x_1; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_16); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); -return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("value is none", 13, 13); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___boxed), 14, 0); +return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2; -x_3 = lean_unsigned_to_nat(16u); -x_4 = lean_unsigned_to_nat(14u); -x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_13; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1; +lean_inc(x_1); +x_15 = lean_mk_array(x_1, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_1, x_16); +lean_dec(x_1); +x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_2, x_15, x_17); +lean_inc(x_4); +x_19 = l_Lean_mkAppN(x_4, x_18); +lean_dec(x_18); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; -x_19 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_18); -if (x_3 == 0) +x_20 = l_Lean_Meta_mkHEq(x_3, x_19, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_20) == 0) { +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; uint8_t x_27; uint8_t x_28; lean_object* x_29; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_4); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_array_mk(x_24); +x_26 = 0; +x_27 = 1; +x_28 = 1; +x_29 = l_Lean_Meta_mkLambdaFVars(x_25, x_21, x_26, x_27, x_26, x_28, x_9, x_10, x_11, x_12, x_22); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_ctor_set(x_13, 0, x_19); -return x_13; -} -else -{ -lean_object* x_20; -lean_free_object(x_13); -x_20 = l_Lean_Meta_mkHEqOfEq(x_19, x_8, x_9, x_10, x_11, x_16); -return x_20; -} +lean_dec(x_25); +return x_29; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_dec(x_13); -x_22 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; -x_23 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_22); -if (x_3 == 0) -{ -lean_object* x_24; +uint8_t x_30; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -return x_24; +lean_dec(x_4); +x_30 = !lean_is_exclusive(x_20); +if (x_30 == 0) +{ +return x_20; } else { -lean_object* x_25; -x_25 = l_Lean_Meta_mkHEqOfEq(x_23, x_8, x_9, x_10, x_11, x_21); -return x_25; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_20, 0); +x_32 = lean_ctor_get(x_20, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_20); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } -else -{ -if (x_3 == 0) -{ -uint8_t x_26; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_26 = !lean_is_exclusive(x_13); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_13, 0); -lean_dec(x_27); -x_28 = lean_ctor_get(x_14, 0); -lean_inc(x_28); -lean_dec(x_14); -lean_ctor_set(x_13, 0, x_28); -return x_13; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1() { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_13, 1); -lean_inc(x_29); -lean_dec(x_13); -x_30 = lean_ctor_get(x_14, 0); -lean_inc(x_30); -lean_dec(x_14); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -return x_31; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Proof.0.Lean.Meta.Grind.mkHCongrProof", 69, 69); +return x_1; } } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_13, 1); -lean_inc(x_32); -lean_dec(x_13); -x_33 = lean_ctor_get(x_14, 0); -lean_inc(x_33); -lean_dec(x_14); -x_34 = l_Lean_Meta_mkHEqOfEq(x_33, x_8, x_9, x_10, x_11, x_32); -return x_34; -} -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1; +x_3 = lean_unsigned_to_nat(135u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -else -{ -uint8_t x_35; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_35 = !lean_is_exclusive(x_13); -if (x_35 == 0) -{ -return x_13; } -else +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3() { +_start: { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_13, 0); -x_37 = lean_ctor_get(x_13, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_13); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("thm.argKinds.size == numArgs\n ", 33, 33); +return x_1; } } +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3; +x_3 = lean_string_append(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5() { _start: { -uint8_t x_12; -x_12 = l_Lean_Expr_isApp(x_1); -if (x_12 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1; +x_3 = lean_unsigned_to_nat(137u); +x_4 = lean_unsigned_to_nat(4u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6() { +_start: { -lean_object* x_13; lean_object* x_14; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_11); -return x_14; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_13 = l_Lean_Expr_getAppFn(x_1); +x_14 = l_Lean_Expr_getAppFn(x_2); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_15); +x_17 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_15); +x_18 = lean_nat_dec_eq(x_17, x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_19 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2; +x_20 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_20; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = l_Lean_Expr_appArg_x21(x_1); -x_16 = l_Lean_Expr_appArg_x21(x_2); -x_17 = l_Lean_Expr_appFn_x21(x_1); -x_18 = l_Lean_Expr_appFn_x21(x_2); +lean_object* x_21; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_16); +lean_inc(x_13); +x_21 = l_Lean_Meta_Grind_mkHCongrWithArity(x_13, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_ctor_get(x_22, 2); +lean_inc(x_24); +x_25 = lean_array_get_size(x_24); +lean_dec(x_24); +x_26 = lean_nat_dec_eq(x_25, x_16); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_27 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5; +x_28 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_28; +} +else +{ +lean_object* x_29; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -8019,130 +7740,158 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_19 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(x_17, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_18); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -if (lean_obj_tag(x_20) == 0) +x_29 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(x_22, x_1, x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +lean_dec(x_16); +lean_dec(x_22); +if (lean_obj_tag(x_29) == 0) { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_13, x_14); +if (x_32 == 0) { -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_19, 1); -x_23 = lean_ctor_get(x_19, 0); -lean_dec(x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); -if (x_24 == 0) +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7; +x_34 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_33, x_10, x_11, x_31); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_13); +x_37 = lean_infer_type(x_13, x_8, x_9, x_10, x_11, x_36); +if (lean_obj_tag(x_37) == 0) { -uint8_t x_25; lean_object* x_26; -lean_free_object(x_19); -x_25 = 0; +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; uint8_t x_42; lean_object* x_43; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___boxed), 13, 3); +lean_closure_set(x_40, 0, x_17); +lean_closure_set(x_40, 1, x_2); +lean_closure_set(x_40, 2, x_1); +x_41 = 0; +x_42 = 0; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_26 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Meta_mkCongrArg(x_17, x_27, x_7, x_8, x_9, x_10, x_28); -if (lean_obj_tag(x_29) == 0) +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_43 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg(x_35, x_41, x_38, x_40, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); +if (lean_obj_tag(x_43) == 0) { -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = 0; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_47 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_13, x_14, x_46, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); +if (lean_obj_tag(x_47) == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_29, 0); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_29, 0, x_32); -return x_29; -} -else +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_50 = l_Lean_Meta_mkEqNDRec(x_44, x_30, x_48, x_8, x_9, x_10, x_11, x_49); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_ctor_get(x_29, 0); -x_34 = lean_ctor_get(x_29, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_29); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_33); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -return x_36; -} +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(x_51, x_3, x_8, x_9, x_10, x_11, x_52); +return x_53; } else { -uint8_t x_37; -x_37 = !lean_is_exclusive(x_29); -if (x_37 == 0) +uint8_t x_54; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_54 = !lean_is_exclusive(x_50); +if (x_54 == 0) { -return x_29; +return x_50; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_29, 0); -x_39 = lean_ctor_get(x_29, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_29); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_50, 0); +x_56 = lean_ctor_get(x_50, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_50); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } } else { -uint8_t x_41; -lean_dec(x_17); +uint8_t x_58; +lean_dec(x_44); +lean_dec(x_30); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -x_41 = !lean_is_exclusive(x_26); -if (x_41 == 0) +x_58 = !lean_is_exclusive(x_47); +if (x_58 == 0) { -return x_26; +return x_47; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_26, 0); -x_43 = lean_ctor_get(x_26, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_26); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_47, 0); +x_60 = lean_ctor_get(x_47, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_47); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; } } } else { -lean_object* x_45; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); +uint8_t x_62; +lean_dec(x_30); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8150,124 +7899,125 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_45 = lean_box(0); -lean_ctor_set(x_19, 0, x_45); -return x_19; -} +x_62 = !lean_is_exclusive(x_43); +if (x_62 == 0) +{ +return x_43; } else { -lean_object* x_46; uint8_t x_47; -x_46 = lean_ctor_get(x_19, 1); -lean_inc(x_46); -lean_dec(x_19); -x_47 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); -if (x_47 == 0) +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_43, 0); +x_64 = lean_ctor_get(x_43, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_43); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +} +} +else { -uint8_t x_48; lean_object* x_49; -x_48 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_49 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_46); -if (lean_obj_tag(x_49) == 0) +uint8_t x_66; +lean_dec(x_35); +lean_dec(x_30); +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_37); +if (x_66 == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_Meta_mkCongrArg(x_17, x_50, x_7, x_8, x_9, x_10, x_51); -if (lean_obj_tag(x_52) == 0) +return x_37; +} +else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - lean_ctor_release(x_52, 1); - x_55 = x_52; -} else { - lean_dec_ref(x_52); - x_55 = lean_box(0); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_37, 0); +x_68 = lean_ctor_get(x_37, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_37); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; } -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_53); -if (lean_is_scalar(x_55)) { - x_57 = lean_alloc_ctor(0, 2, 0); -} else { - x_57 = x_55; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -return x_57; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_58 = lean_ctor_get(x_52, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_52, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - lean_ctor_release(x_52, 1); - x_60 = x_52; -} else { - lean_dec_ref(x_52); - x_60 = lean_box(0); -} -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(1, 2, 0); -} else { - x_61 = x_60; -} -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_59); -return x_61; +lean_object* x_70; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_70 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqOfHEqIfNeeded(x_30, x_3, x_8, x_9, x_10, x_11, x_31); +return x_70; } } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +uint8_t x_71; lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_62 = lean_ctor_get(x_49, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_49, 1); -lean_inc(x_63); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_64 = x_49; -} else { - lean_dec_ref(x_49); - x_64 = lean_box(0); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_71 = !lean_is_exclusive(x_29); +if (x_71 == 0) +{ +return x_29; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_29, 0); +x_73 = lean_ctor_get(x_29, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_29); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; } -if (lean_is_scalar(x_64)) { - x_65 = lean_alloc_ctor(1, 2, 0); -} else { - x_65 = x_64; } -lean_ctor_set(x_65, 0, x_62); -lean_ctor_set(x_65, 1, x_63); -return x_65; } } else { -lean_object* x_66; lean_object* x_67; +uint8_t x_75; lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8275,387 +8025,1554 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_66 = lean_box(0); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_46); -return x_67; +lean_dec(x_2); +lean_dec(x_1); +x_75 = !lean_is_exclusive(x_21); +if (x_75 == 0) +{ +return x_21; } +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_21, 0); +x_77 = lean_ctor_get(x_21, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_21); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } -else +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1() { +_start: { -lean_object* x_68; uint8_t x_69; -lean_dec(x_17); -x_68 = lean_ctor_get(x_19, 1); -lean_inc(x_68); -lean_dec(x_19); -x_69 = !lean_is_exclusive(x_20); -if (x_69 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2() { +_start: { -lean_object* x_70; uint8_t x_71; -x_70 = lean_ctor_get(x_20, 0); -x_71 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); -if (x_71 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3() { +_start: { -uint8_t x_72; lean_object* x_73; -x_72 = 0; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("value is none", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -x_73 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_72, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); -if (lean_obj_tag(x_73) == 0) +x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(x_1, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_73, 1); -lean_inc(x_75); -lean_dec(x_73); -x_76 = l_Lean_Meta_mkCongr(x_70, x_74, x_7, x_8, x_9, x_10, x_75); -if (lean_obj_tag(x_76) == 0) +lean_object* x_14; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_77; -x_77 = !lean_is_exclusive(x_76); -if (x_77 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; +x_19 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_18); +if (x_3 == 0) +{ +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; +lean_free_object(x_13); +x_20 = l_Lean_Meta_mkHEqOfEq(x_19, x_8, x_9, x_10, x_11, x_16); +return x_20; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_dec(x_13); +x_22 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4; +x_23 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_22); +if (x_3 == 0) +{ +lean_object* x_24; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +return x_24; +} +else +{ +lean_object* x_25; +x_25 = l_Lean_Meta_mkHEqOfEq(x_23, x_8, x_9, x_10, x_11, x_21); +return x_25; +} +} +} +else +{ +if (x_3 == 0) +{ +uint8_t x_26; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_13, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_14, 0); +lean_inc(x_28); +lean_dec(x_14); +lean_ctor_set(x_13, 0, x_28); +return x_13; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_13, 1); +lean_inc(x_29); +lean_dec(x_13); +x_30 = lean_ctor_get(x_14, 0); +lean_inc(x_30); +lean_dec(x_14); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_13, 1); +lean_inc(x_32); +lean_dec(x_13); +x_33 = lean_ctor_get(x_14, 0); +lean_inc(x_33); +lean_dec(x_14); +x_34 = l_Lean_Meta_mkHEqOfEq(x_33, x_8, x_9, x_10, x_11, x_32); +return x_34; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_35 = !lean_is_exclusive(x_13); +if (x_35 == 0) +{ +return x_13; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_13, 0); +x_37 = lean_ctor_get(x_13, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_13); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = l_Lean_Expr_isApp(x_1); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_11); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = l_Lean_Expr_appArg_x21(x_1); +x_16 = l_Lean_Expr_appArg_x21(x_2); +x_17 = l_Lean_Expr_appFn_x21(x_1); +x_18 = l_Lean_Expr_appFn_x21(x_2); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_19 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_loop(x_17, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_19, 1); +x_23 = lean_ctor_get(x_19, 0); +lean_dec(x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); +if (x_24 == 0) +{ +uint8_t x_25; lean_object* x_26; +lean_free_object(x_19); +x_25 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_26 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l_Lean_Meta_mkCongrArg(x_17, x_27, x_7, x_8, x_9, x_10, x_28); +if (lean_obj_tag(x_29) == 0) +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_29, 0, x_32); +return x_29; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_29, 0); +x_34 = lean_ctor_get(x_29, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_29); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_33); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +return x_36; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_29); +if (x_37 == 0) +{ +return x_29; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_29, 0); +x_39 = lean_ctor_get(x_29, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_29); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +else +{ +uint8_t x_41; +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_41 = !lean_is_exclusive(x_26); +if (x_41 == 0) +{ +return x_26; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_26); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +else +{ +lean_object* x_45; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_45 = lean_box(0); +lean_ctor_set(x_19, 0, x_45); +return x_19; +} +} +else +{ +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_19, 1); +lean_inc(x_46); +lean_dec(x_19); +x_47 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); +if (x_47 == 0) +{ +uint8_t x_48; lean_object* x_49; +x_48 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_49 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_46); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Lean_Meta_mkCongrArg(x_17, x_50, x_7, x_8, x_9, x_10, x_51); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_55 = x_52; +} else { + lean_dec_ref(x_52); + x_55 = lean_box(0); +} +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_53); +if (lean_is_scalar(x_55)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_55; +} +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +return x_57; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_52, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_52, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_60 = x_52; +} else { + lean_dec_ref(x_52); + x_60 = lean_box(0); +} +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); +} else { + x_61 = x_60; +} +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; +} +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_62 = lean_ctor_get(x_49, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_49, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_64 = x_49; +} else { + lean_dec_ref(x_49); + x_64 = lean_box(0); +} +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(1, 2, 0); +} else { + x_65 = x_64; +} +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_63); +return x_65; +} +} +else +{ +lean_object* x_66; lean_object* x_67; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_66 = lean_box(0); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_46); +return x_67; +} +} +} +else +{ +lean_object* x_68; uint8_t x_69; +lean_dec(x_17); +x_68 = lean_ctor_get(x_19, 1); +lean_inc(x_68); +lean_dec(x_19); +x_69 = !lean_is_exclusive(x_20); +if (x_69 == 0) +{ +lean_object* x_70; uint8_t x_71; +x_70 = lean_ctor_get(x_20, 0); +x_71 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); +if (x_71 == 0) +{ +uint8_t x_72; lean_object* x_73; +x_72 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_73 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_72, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_mkCongr(x_70, x_74, x_7, x_8, x_9, x_10, x_75); +if (lean_obj_tag(x_76) == 0) +{ +uint8_t x_77; +x_77 = !lean_is_exclusive(x_76); +if (x_77 == 0) +{ +lean_object* x_78; +x_78 = lean_ctor_get(x_76, 0); +lean_ctor_set(x_20, 0, x_78); +lean_ctor_set(x_76, 0, x_20); +return x_76; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_76, 0); +x_80 = lean_ctor_get(x_76, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_76); +lean_ctor_set(x_20, 0, x_79); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_20); +lean_ctor_set(x_81, 1, x_80); +return x_81; +} +} +else +{ +uint8_t x_82; +lean_free_object(x_20); +x_82 = !lean_is_exclusive(x_76); +if (x_82 == 0) { -lean_object* x_78; -x_78 = lean_ctor_get(x_76, 0); -lean_ctor_set(x_20, 0, x_78); -lean_ctor_set(x_76, 0, x_20); return x_76; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_76, 0); -x_80 = lean_ctor_get(x_76, 1); +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_76, 0); +x_84 = lean_ctor_get(x_76, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_76); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +return x_85; +} +} +} +else +{ +uint8_t x_86; +lean_free_object(x_20); +lean_dec(x_70); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_86 = !lean_is_exclusive(x_73); +if (x_86 == 0) +{ +return x_73; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_73, 0); +x_88 = lean_ctor_get(x_73, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_73); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; +} +} +} +else +{ +lean_object* x_90; +lean_dec(x_16); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_90 = l_Lean_Meta_mkCongrFun(x_70, x_15, x_7, x_8, x_9, x_10, x_68); +if (lean_obj_tag(x_90) == 0) +{ +uint8_t x_91; +x_91 = !lean_is_exclusive(x_90); +if (x_91 == 0) +{ +lean_object* x_92; +x_92 = lean_ctor_get(x_90, 0); +lean_ctor_set(x_20, 0, x_92); +lean_ctor_set(x_90, 0, x_20); +return x_90; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_90, 0); +x_94 = lean_ctor_get(x_90, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_90); +lean_ctor_set(x_20, 0, x_93); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_20); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +} +else +{ +uint8_t x_96; +lean_free_object(x_20); +x_96 = !lean_is_exclusive(x_90); +if (x_96 == 0) +{ +return x_90; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_90, 0); +x_98 = lean_ctor_get(x_90, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_90); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; +} +} +} +} +else +{ +lean_object* x_100; uint8_t x_101; +x_100 = lean_ctor_get(x_20, 0); +lean_inc(x_100); +lean_dec(x_20); +x_101 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); +if (x_101 == 0) +{ +uint8_t x_102; lean_object* x_103; +x_102 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_103 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); +x_106 = l_Lean_Meta_mkCongr(x_100, x_104, x_7, x_8, x_9, x_10, x_105); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_109 = x_106; +} else { + lean_dec_ref(x_106); + x_109 = lean_box(0); +} +x_110 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_110, 0, x_107); +if (lean_is_scalar(x_109)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_109; +} +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_108); +return x_111; +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_112 = lean_ctor_get(x_106, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_106, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_114 = x_106; +} else { + lean_dec_ref(x_106); + x_114 = lean_box(0); +} +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(1, 2, 0); +} else { + x_115 = x_114; +} +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_113); +return x_115; +} +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_100); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_116 = lean_ctor_get(x_103, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_103, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_118 = x_103; +} else { + lean_dec_ref(x_103); + x_118 = lean_box(0); +} +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); +} else { + x_119 = x_118; +} +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; +} +} +else +{ +lean_object* x_120; +lean_dec(x_16); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_120 = l_Lean_Meta_mkCongrFun(x_100, x_15, x_7, x_8, x_9, x_10, x_68); +if (lean_obj_tag(x_120) == 0) +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_123 = x_120; +} else { + lean_dec_ref(x_120); + x_123 = lean_box(0); +} +x_124 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_124, 0, x_121); +if (lean_is_scalar(x_123)) { + x_125 = lean_alloc_ctor(0, 2, 0); +} else { + x_125 = x_123; +} +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_122); +return x_125; +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_120, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_120, 1); +lean_inc(x_127); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_128 = x_120; +} else { + lean_dec_ref(x_120); + x_128 = lean_box(0); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(1, 2, 0); +} else { + x_129 = x_128; +} +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_127); +return x_129; +} +} +} +} +} +else +{ +uint8_t x_130; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_130 = !lean_is_exclusive(x_19); +if (x_130 == 0) +{ +return x_19; +} +else +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_19, 0); +x_132 = lean_ctor_get(x_19, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_19); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +return x_133; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hasSameRoot enodes a₁ b₂ && hasSameRoot enodes b₁ a₂\n ", 67, 59); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Proof.0.Lean.Meta.Grind.mkEqCongrProof", 70, 70); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(174u); +x_4 = lean_unsigned_to_nat(6u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_congr'", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; +x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eq_congr", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; +x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +_start: +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_57; +x_20 = lean_st_ref_get(x_11, x_19); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Expr_constLevels_x21(x_1); +x_57 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_8); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_58 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_5, x_6, x_7, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +return x_58; +} +else +{ +uint8_t x_59; +lean_dec(x_6); +lean_dec(x_5); +lean_inc(x_23); +x_59 = l_Lean_Meta_Grind_hasSameRoot(x_23, x_2, x_9); +if (x_59 == 0) +{ +lean_object* x_60; +x_60 = lean_box(0); +x_25 = x_60; +goto block_56; +} +else +{ +uint8_t x_61; +lean_inc(x_23); +x_61 = l_Lean_Meta_Grind_hasSameRoot(x_23, x_3, x_10); +if (x_61 == 0) +{ +lean_object* x_62; +x_62 = lean_box(0); +x_25 = x_62; +goto block_56; +} +else +{ +uint8_t x_63; lean_object* x_64; +lean_dec(x_23); +x_63 = 0; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_9); +lean_inc(x_2); +x_64 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_2, x_9, x_63, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +lean_inc(x_10); +lean_inc(x_3); +x_67 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_3, x_10, x_63, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_66); +if (lean_obj_tag(x_67) == 0) +{ +uint8_t x_68; +x_68 = !lean_is_exclusive(x_67); +if (x_68 == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_67, 0); +x_70 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8; +x_71 = l_Lean_Expr_const___override(x_70, x_24); +x_72 = l_Lean_mkApp7(x_71, x_4, x_2, x_3, x_9, x_10, x_65, x_69); +lean_ctor_set(x_67, 0, x_72); +return x_67; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_73 = lean_ctor_get(x_67, 0); +x_74 = lean_ctor_get(x_67, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_67); +x_75 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8; +x_76 = l_Lean_Expr_const___override(x_75, x_24); +x_77 = l_Lean_mkApp7(x_76, x_4, x_2, x_3, x_9, x_10, x_65, x_73); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_74); +return x_78; +} +} +else +{ +uint8_t x_79; +lean_dec(x_65); +lean_dec(x_24); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_79 = !lean_is_exclusive(x_67); +if (x_79 == 0) +{ +return x_67; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_67, 0); +x_81 = lean_ctor_get(x_67, 1); +lean_inc(x_81); lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_76); -lean_ctor_set(x_20, 0, x_79); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_20); -lean_ctor_set(x_81, 1, x_80); -return x_81; +lean_dec(x_67); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; +} +} +} +else +{ +uint8_t x_83; +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_83 = !lean_is_exclusive(x_64); +if (x_83 == 0) +{ +return x_64; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_64, 0); +x_85 = lean_ctor_get(x_64, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_64); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; +} +} +} +} } +block_56: +{ +uint8_t x_26; +lean_dec(x_25); +lean_inc(x_23); +x_26 = l_Lean_Meta_Grind_hasSameRoot(x_23, x_2, x_10); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_27 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4; +x_28 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_27, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +return x_28; } else { -uint8_t x_82; -lean_free_object(x_20); -x_82 = !lean_is_exclusive(x_76); -if (x_82 == 0) +uint8_t x_29; +x_29 = l_Lean_Meta_Grind_hasSameRoot(x_23, x_3, x_9); +if (x_29 == 0) { -return x_76; +lean_object* x_30; lean_object* x_31; +lean_dec(x_24); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_30 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4; +x_31 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +return x_31; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_76, 0); -x_84 = lean_ctor_get(x_76, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_76); -x_85 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -return x_85; +uint8_t x_32; lean_object* x_33; +x_32 = 0; +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_33 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_2, x_10, x_32, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +lean_inc(x_9); +lean_inc(x_3); +x_36 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_3, x_9, x_32, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_35); +if (lean_obj_tag(x_36) == 0) +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_36, 0); +x_39 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6; +x_40 = l_Lean_Expr_const___override(x_39, x_24); +x_41 = l_Lean_mkApp7(x_40, x_4, x_2, x_3, x_9, x_10, x_34, x_38); +lean_ctor_set(x_36, 0, x_41); +return x_36; } +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_42 = lean_ctor_get(x_36, 0); +x_43 = lean_ctor_get(x_36, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_36); +x_44 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6; +x_45 = l_Lean_Expr_const___override(x_44, x_24); +x_46 = l_Lean_mkApp7(x_45, x_4, x_2, x_3, x_9, x_10, x_34, x_42); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_43); +return x_47; } } else { -uint8_t x_86; -lean_free_object(x_20); -lean_dec(x_70); +uint8_t x_48; +lean_dec(x_34); +lean_dec(x_24); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_86 = !lean_is_exclusive(x_73); -if (x_86 == 0) +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_48 = !lean_is_exclusive(x_36); +if (x_48 == 0) { -return x_73; +return x_36; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_73, 0); -x_88 = lean_ctor_get(x_73, 1); -lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_73); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_36, 0); +x_50 = lean_ctor_get(x_36, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_36); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } else { -lean_object* x_90; +uint8_t x_52; +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_17); lean_dec(x_16); -lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_4); lean_dec(x_3); -x_90 = l_Lean_Meta_mkCongrFun(x_70, x_15, x_7, x_8, x_9, x_10, x_68); -if (lean_obj_tag(x_90) == 0) -{ -uint8_t x_91; -x_91 = !lean_is_exclusive(x_90); -if (x_91 == 0) +lean_dec(x_2); +x_52 = !lean_is_exclusive(x_33); +if (x_52 == 0) { -lean_object* x_92; -x_92 = lean_ctor_get(x_90, 0); -lean_ctor_set(x_20, 0, x_92); -lean_ctor_set(x_90, 0, x_20); -return x_90; +return x_33; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_90, 0); -x_94 = lean_ctor_get(x_90, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_90); -lean_ctor_set(x_20, 0, x_93); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_20); -lean_ctor_set(x_95, 1, x_94); -return x_95; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_33, 0); +x_54 = lean_ctor_get(x_33, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_33); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } -else -{ -uint8_t x_96; -lean_free_object(x_20); -x_96 = !lean_is_exclusive(x_90); -if (x_96 == 0) -{ -return x_90; } -else -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_90, 0); -x_98 = lean_ctor_get(x_90, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_90); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -return x_99; } } } } -else -{ -lean_object* x_100; uint8_t x_101; -x_100 = lean_ctor_get(x_20, 0); -lean_inc(x_100); -lean_dec(x_20); -x_101 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_16); -if (x_101 == 0) -{ -uint8_t x_102; lean_object* x_103; -x_102 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_103 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore(x_15, x_16, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); -if (lean_obj_tag(x_103) == 0) -{ -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_103, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_103, 1); -lean_inc(x_105); -lean_dec(x_103); -x_106 = l_Lean_Meta_mkCongr(x_100, x_104, x_7, x_8, x_9, x_10, x_105); -if (lean_obj_tag(x_106) == 0) +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1() { +_start: { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_109 = x_106; -} else { - lean_dec_ref(x_106); - x_109 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(166u); +x_4 = lean_unsigned_to_nat(34u); +x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -x_110 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_110, 0, x_107); -if (lean_is_scalar(x_109)) { - x_111 = lean_alloc_ctor(0, 2, 0); -} else { - x_111 = x_109; } -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_108); -return x_111; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1; +x_12 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } -else +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1() { +_start: { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_112 = lean_ctor_get(x_106, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_106, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_114 = x_106; -} else { - lean_dec_ref(x_106); - x_114 = lean_box(0); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___boxed), 10, 0); +return x_1; } -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(1, 2, 0); -} else { - x_115 = x_114; } -lean_ctor_set(x_115, 0, x_112); -lean_ctor_set(x_115, 1, x_113); -return x_115; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1; +lean_inc(x_1); +x_18 = l_Lean_Expr_cleanupAnnotations(x_1); +x_19 = l_Lean_Expr_isApp(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_18); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_20 = lean_box(0); +x_21 = lean_apply_10(x_17, x_20, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_21; } +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = l_Lean_Expr_appArg(x_18, lean_box(0)); +x_23 = l_Lean_Expr_appFnCleanup(x_18, lean_box(0)); +x_24 = l_Lean_Expr_isApp(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_25 = lean_box(0); +x_26 = lean_apply_10(x_17, x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_26; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -lean_dec(x_100); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = l_Lean_Expr_appArg(x_23, lean_box(0)); +x_28 = l_Lean_Expr_appFnCleanup(x_23, lean_box(0)); +x_29 = l_Lean_Expr_isApp(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_28); +lean_dec(x_27); +lean_dec(x_22); lean_dec(x_7); -x_116 = lean_ctor_get(x_103, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_103, 1); -lean_inc(x_117); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_118 = x_103; -} else { - lean_dec_ref(x_103); - x_118 = lean_box(0); -} -if (lean_is_scalar(x_118)) { - x_119 = lean_alloc_ctor(1, 2, 0); -} else { - x_119 = x_118; -} -lean_ctor_set(x_119, 0, x_116); -lean_ctor_set(x_119, 1, x_117); -return x_119; -} +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_30 = lean_box(0); +x_31 = lean_apply_10(x_17, x_30, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_31; } else { -lean_object* x_120; -lean_dec(x_16); +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = l_Lean_Expr_appArg(x_28, lean_box(0)); +x_33 = l_Lean_Expr_appFnCleanup(x_28, lean_box(0)); +x_34 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; +x_35 = l_Lean_Expr_isConstOf(x_33, x_34); +lean_dec(x_33); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_32); +lean_dec(x_27); +lean_dec(x_22); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_120 = l_Lean_Meta_mkCongrFun(x_100, x_15, x_7, x_8, x_9, x_10, x_68); -if (lean_obj_tag(x_120) == 0) +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_box(0); +x_37 = lean_apply_10(x_17, x_36, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_37; +} +else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_121 = lean_ctor_get(x_120, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 1); -lean_inc(x_122); -if (lean_is_exclusive(x_120)) { - lean_ctor_release(x_120, 0); - lean_ctor_release(x_120, 1); - x_123 = x_120; -} else { - lean_dec_ref(x_120); - x_123 = lean_box(0); +lean_object* x_38; +x_38 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1(x_4, x_6, x_7, x_5, x_2, x_1, x_3, x_32, x_27, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_32); +return x_38; } -x_124 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_124, 0, x_121); -if (lean_is_scalar(x_123)) { - x_125 = lean_alloc_ctor(0, 2, 0); -} else { - x_125 = x_123; } -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_122); -return x_125; } -else +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1() { +_start: { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_126 = lean_ctor_get(x_120, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_120, 1); -lean_inc(x_127); -if (lean_is_exclusive(x_120)) { - lean_ctor_release(x_120, 0); - lean_ctor_release(x_120, 1); - x_128 = x_120; -} else { - lean_dec_ref(x_120); - x_128 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3; +x_3 = lean_unsigned_to_nat(165u); +x_4 = lean_unsigned_to_nat(36u); +x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } -if (lean_is_scalar(x_128)) { - x_129 = lean_alloc_ctor(1, 2, 0); -} else { - x_129 = x_128; } -lean_ctor_set(x_129, 0, x_126); -lean_ctor_set(x_129, 1, x_127); -return x_129; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1; +x_12 = l_panic___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__9(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___boxed), 10, 0); +return x_1; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1; +lean_inc(x_1); +x_14 = l_Lean_Expr_cleanupAnnotations(x_1); +x_15 = l_Lean_Expr_isApp(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_14); +lean_dec(x_2); +lean_dec(x_1); +x_16 = lean_box(0); +x_17 = lean_apply_10(x_13, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_17; } +else +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Lean_Expr_appArg(x_14, lean_box(0)); +x_19 = l_Lean_Expr_appFnCleanup(x_14, lean_box(0)); +x_20 = l_Lean_Expr_isApp(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +x_22 = lean_apply_10(x_13, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_22; } +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = l_Lean_Expr_appArg(x_19, lean_box(0)); +x_24 = l_Lean_Expr_appFnCleanup(x_19, lean_box(0)); +x_25 = l_Lean_Expr_isApp(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_18); +lean_dec(x_2); +lean_dec(x_1); +x_26 = lean_box(0); +x_27 = lean_apply_10(x_13, x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_27; } else { -uint8_t x_130; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_130 = !lean_is_exclusive(x_19); -if (x_130 == 0) +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = l_Lean_Expr_appArg(x_24, lean_box(0)); +x_29 = l_Lean_Expr_appFnCleanup(x_24, lean_box(0)); +x_30 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_isEqProof___closed__2; +x_31 = l_Lean_Expr_isConstOf(x_29, x_30); +if (x_31 == 0) { -return x_19; +lean_object* x_32; lean_object* x_33; +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_23); +lean_dec(x_18); +lean_dec(x_2); +lean_dec(x_1); +x_32 = lean_box(0); +x_33 = lean_apply_10(x_13, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_33; } else { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_19, 0); -x_132 = lean_ctor_get(x_19, 1); -lean_inc(x_132); -lean_inc(x_131); -lean_dec(x_19); -x_133 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -return x_133; +lean_object* x_34; +x_34 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3(x_2, x_1, x_3, x_29, x_28, x_23, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_29); +return x_34; +} } } } @@ -8673,8 +9590,8 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10; -x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11; +x_1 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__5; +x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__6; x_3 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -8769,7 +9686,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(181u); +x_3 = lean_unsigned_to_nat(203u); x_4 = lean_unsigned_to_nat(35u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -8782,7 +9699,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofTo___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(182u); +x_3 = lean_unsigned_to_nat(204u); x_4 = lean_unsigned_to_nat(29u); x_5 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -9133,11 +10050,11 @@ return x_17; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; -x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof_loop(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof_loop(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -9201,7 +10118,17 @@ x_17 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_realizeEqProo return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_3); +lean_dec(x_3); +x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { uint8_t x_15; uint8_t x_16; lean_object* x_17; @@ -9209,29 +10136,29 @@ x_15 = lean_unbox(x_2); lean_dec(x_2); x_16 = lean_unbox(x_5); lean_dec(x_5); -x_17 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___spec__1___rarg(x_1, x_15, x_3, x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_17 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___spec__1___rarg(x_1, x_15, x_3, x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_13; -x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_14; +x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_13; +return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; lean_object* x_14; x_13 = lean_unbox(x_3); lean_dec(x_3); -x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_14; } } @@ -9257,6 +10184,76 @@ lean_dec(x_1); return x_12; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +_start: +{ +uint8_t x_20; lean_object* x_21; +x_20 = lean_unbox(x_7); +lean_dec(x_7); +x_21 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_20, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_8); +lean_dec(x_1); +return x_21; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_3); +lean_dec(x_3); +x_18 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3(x_1, x_2, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_4); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_3); +lean_dec(x_3); +x_14 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -9296,7 +10293,7 @@ static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Proof._hyg.2379.0 )\n ", 60, 60); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Meta.Tactic.Grind.Proof._hyg.3014.0 )\n ", 60, 60); return x_1; } } @@ -9324,7 +10321,7 @@ static lean_object* _init_l_Lean_Meta_Grind_mkEqProofImpl___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Loop_forIn_loop___at___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_findCommon___spec__7___lambda__1___closed__1; x_2 = l_Lean_Meta_Grind_mkEqProofImpl___closed__3; -x_3 = lean_unsigned_to_nat(231u); +x_3 = lean_unsigned_to_nat(253u); x_4 = lean_unsigned_to_nat(2u); x_5 = l_Lean_Meta_Grind_mkEqProofImpl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -9415,14 +10412,14 @@ x_13 = l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqProofCore return x_13; } } -lean_object* initialize_Lean_Meta_Sorry(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Grind_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Proof(uint8_t builtin, lean_object* w) { lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Lean_Meta_Sorry(builtin, lean_io_mk_world()); +res = initialize_Init_Grind_Lemmas(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); @@ -9476,8 +10473,6 @@ l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__2); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__3(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkProofFrom___lambda__1___closed__3); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___lambda__1___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__2(); @@ -9494,16 +10489,22 @@ l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___close lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__7); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__8); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__9); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__10); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__11); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__12); -l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrProof___closed__13); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__2); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__3); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__4); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__5); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__6); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkHCongrProof___closed__7); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__2(); @@ -9512,6 +10513,30 @@ l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof_ lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__3); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkCongrDefaultProof___closed__4); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__2); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__3); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__4); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__5); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__6); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__7); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__1___closed__8); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___lambda__4___closed__1); +l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkEqCongrProof___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__1); l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Proof_0__Lean_Meta_Grind_mkNestedProofCongr___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c index 80e456e2dbda..7f380dd9a3ac 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Propagate.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Propagate -// Imports: Init.Grind Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.PropagatorAttr +// Imports: Init.Grind Lean.Meta.Tactic.Grind.Proof Lean.Meta.Tactic.Grind.PropagatorAttr Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.Internalize #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,10 +15,12 @@ extern "C" { #endif static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateEqMatchDown___closed__2; static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__2; -static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7; +lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_383____closed__1; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__11; static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__4; @@ -26,133 +28,171 @@ static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__5; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_383_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__10; static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__2; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058_(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateIte___closed__1; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__5; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__2; static lean_object* l_Lean_Meta_Grind_propagateNotUp___closed__2; +lean_object* l_Lean_mkApp8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6; lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__5; +static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateOrUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_985____closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__3; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6; static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__8; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1; static lean_object* l_Lean_Meta_Grind_propagateHEqDown___closed__1; lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__1; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_604____closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqMatchDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___closed__2; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateEqMatchDown___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateHEqDown___closed__2; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706_(lean_object*); +lean_object* l_Lean_Meta_Grind_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkApp6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__4; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__7; static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__2; +lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__7; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__12; static lean_object* l_Lean_Meta_Grind_propagateAndUp___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2; lean_object* l_Lean_Meta_Grind_closeGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__4; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__8; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__6; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__2; -static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8; static lean_object* l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateOrDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1206_(lean_object*); -static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateOrDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1206____closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__11; static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__4; lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__4; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5; static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__1; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__8; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7; static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__3; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__10; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1; static lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__5; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148_(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1; +static lean_object* l_Lean_Meta_Grind_propagateDIte___closed__2; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__3; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__13; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateOrUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_985_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287_(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateIte___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateAndDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_604_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__4; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530_(lean_object*); lean_object* lean_grind_mk_eq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___closed__1; lean_object* lean_grind_mk_heq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_constLevels_x21(lean_object*); +static lean_object* l_Lean_Meta_Grind_propagateDIte___closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1; +static lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4; +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___closed__1; static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__6; static lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1___closed__9; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__2; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_isEqv(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347_(lean_object*); lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426_(lean_object*); +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndDown___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqMatchDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_pushEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrDown___lambda__1___closed__5; +static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1; +static lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3; lean_object* l_Lean_Meta_Grind_isEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__12; static lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__3; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546_(lean_object*); static lean_object* l_Lean_Meta_Grind_propagateOrUp___closed__2; -static lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1; static lean_object* l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__9; static lean_object* l_Lean_Meta_Grind_propagateNotUp___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateOrUp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_propagateAndUp___closed__3; +static lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6; static lean_object* l_Lean_Meta_Grind_propagateEqUp___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078_(lean_object*); static lean_object* _init_l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1() { _start: { @@ -2554,7 +2594,7 @@ static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("not_eq_of_eq_true", 17, 17); +x_1 = lean_mk_string_unchecked("false_of_not_eq_self", 20, 20); return x_1; } } @@ -2583,7 +2623,7 @@ static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("not_eq_of_eq_false", 18, 18); +x_1 = lean_mk_string_unchecked("not_eq_of_eq_true", 17, 17); return x_1; } } @@ -2608,6 +2648,35 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("not_eq_of_eq_false", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { @@ -2638,7 +2707,23 @@ x_18 = lean_unbox(x_17); lean_dec(x_17); if (x_18 == 0) { -uint8_t x_19; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +lean_inc(x_2); +lean_inc(x_1); +x_20 = l_Lean_Meta_Grind_isEqv(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_unbox(x_21); +lean_dec(x_21); +if (x_22 == 0) +{ +uint8_t x_23; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -2649,35 +2734,35 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) +x_23 = !lean_is_exclusive(x_20); +if (x_23 == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_16, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_16, 0, x_21); -return x_16; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_20, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_20, 0, x_25); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_16, 1); -lean_inc(x_22); -lean_dec(x_16); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_dec(x_20); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; } } else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_dec(x_16); +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_20, 1); +lean_inc(x_29); +lean_dec(x_20); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -2687,31 +2772,23 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Meta_Grind_mkEqTrueProof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -if (lean_obj_tag(x_26) == 0) +x_30 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_29); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__3; -x_30 = l_Lean_mkAppB(x_29, x_2, x_27); -x_31 = l_Lean_Meta_Grind_pushEqFalse(x_1, x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_31; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__3; +x_34 = l_Lean_mkAppB(x_33, x_2, x_31); +x_35 = l_Lean_Meta_Grind_closeGoal(x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +return x_35; } else { -uint8_t x_32; +uint8_t x_36; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -2721,31 +2798,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_32 = !lean_is_exclusive(x_26); -if (x_32 == 0) +x_36 = !lean_is_exclusive(x_30); +if (x_36 == 0) { -return x_26; +return x_30; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_26, 0); -x_34 = lean_ctor_get(x_26, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_26); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_30, 0); +x_38 = lean_ctor_get(x_30, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_30); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } } } else { -uint8_t x_36; +uint8_t x_40; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -2756,32 +2832,32 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_16); -if (x_36 == 0) +x_40 = !lean_is_exclusive(x_20); +if (x_40 == 0) { -return x_16; +return x_20; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_16, 0); -x_38 = lean_ctor_get(x_16, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_16); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_20, 0); +x_42 = lean_ctor_get(x_20, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_20); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } } else { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_12, 1); -lean_inc(x_40); -lean_dec(x_12); +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_16, 1); +lean_inc(x_44); +lean_dec(x_16); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -2791,31 +2867,18 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_41 = l_Lean_Meta_Grind_mkEqFalseProof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_40); -if (lean_obj_tag(x_41) == 0) -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6; -x_45 = l_Lean_mkAppB(x_44, x_2, x_42); -x_46 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_43); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_46; -} -else +x_45 = l_Lean_Meta_Grind_mkEqTrueProof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_44); +if (lean_obj_tag(x_45) == 0) { -uint8_t x_47; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6; +x_49 = l_Lean_mkAppB(x_48, x_2, x_46); +x_50 = l_Lean_Meta_Grind_pushEqFalse(x_1, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -2824,29 +2887,8 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_41); -if (x_47 == 0) -{ -return x_41; -} -else -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_41, 0); -x_49 = lean_ctor_get(x_41, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_41); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); return x_50; } -} -} -} else { uint8_t x_51; @@ -2860,19 +2902,19 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_12); +x_51 = !lean_is_exclusive(x_45); if (x_51 == 0) { -return x_12; +return x_45; } else { lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_12, 0); -x_53 = lean_ctor_get(x_12, 1); +x_52 = lean_ctor_get(x_45, 0); +x_53 = lean_ctor_get(x_45, 1); lean_inc(x_53); lean_inc(x_52); -lean_dec(x_12); +lean_dec(x_45); x_54 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_54, 0, x_52); lean_ctor_set(x_54, 1, x_53); @@ -2881,35 +2923,10 @@ return x_54; } } } -static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Not", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateNotUp___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; uint8_t x_12; -lean_inc(x_1); -x_11 = l_Lean_Expr_cleanupAnnotations(x_1); -x_12 = l_Lean_Expr_isApp(x_11); -if (x_12 == 0) +else { -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); +uint8_t x_55; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2919,25 +2936,188 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_10); -return x_14; +x_55 = !lean_is_exclusive(x_16); +if (x_55 == 0) +{ +return x_16; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_Expr_appArg(x_11, lean_box(0)); -x_16 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); -x_17 = l_Lean_Meta_Grind_propagateNotUp___closed__2; -x_18 = l_Lean_Expr_isConstOf(x_16, x_17); +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_16, 0); +x_57 = lean_ctor_get(x_16, 1); +lean_inc(x_57); +lean_inc(x_56); lean_dec(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_15); -lean_dec(x_9); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_12, 1); +lean_inc(x_59); +lean_dec(x_12); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_60 = l_Lean_Meta_Grind_mkEqFalseProof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_59); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9; +x_64 = l_Lean_mkAppB(x_63, x_2, x_61); +x_65 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_64, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_62); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_65; +} +else +{ +uint8_t x_66; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_60); +if (x_66 == 0) +{ +return x_60; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_60, 0); +x_68 = lean_ctor_get(x_60, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_60); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} +} +} +else +{ +uint8_t x_70; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_70 = !lean_is_exclusive(x_12); +if (x_70 == 0) +{ +return x_12; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_12, 0); +x_72 = lean_ctor_get(x_12, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_12); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +return x_73; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Not", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateNotUp___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateNotUp___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +lean_inc(x_1); +x_11 = l_Lean_Expr_cleanupAnnotations(x_1); +x_12 = l_Lean_Expr_isApp(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_10); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_Expr_appArg(x_11, lean_box(0)); +x_16 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); +x_17 = l_Lean_Meta_Grind_propagateNotUp___closed__2; +x_18 = l_Lean_Expr_isConstOf(x_16, x_17); +lean_dec(x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_15); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2961,7 +3141,7 @@ return x_21; } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1() { _start: { lean_object* x_1; @@ -2969,13 +3149,13 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateNotUp), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_Grind_propagateNotUp___closed__2; x_3 = 1; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } @@ -2984,7 +3164,7 @@ static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___close _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("false_of_not_eq_self", 20, 20); +x_1 = lean_mk_string_unchecked("eq_false_of_not_eq_true", 23, 23); return x_1; } } @@ -3013,7 +3193,7 @@ static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___close _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("eq_false_of_not_eq_true", 23, 23); +x_1 = lean_mk_string_unchecked("eq_true_of_not_eq_false", 23, 23); return x_1; } } @@ -3038,35 +3218,6 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("eq_true_of_not_eq_false", 23, 23); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; -x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; -x_3 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateNotDown___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { @@ -3171,13 +3322,9 @@ lean_inc(x_31); x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec(x_30); -x_33 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__3; +x_33 = l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__3; x_34 = l_Lean_mkAppB(x_33, x_2, x_31); x_35 = l_Lean_Meta_Grind_closeGoal(x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); return x_35; } else @@ -3269,7 +3416,7 @@ lean_inc(x_46); x_47 = lean_ctor_get(x_45, 1); lean_inc(x_47); lean_dec(x_45); -x_48 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6; +x_48 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__3; lean_inc(x_2); x_49 = l_Lean_mkAppB(x_48, x_2, x_46); x_50 = l_Lean_Meta_Grind_pushEqFalse(x_2, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_47); @@ -3372,7 +3519,7 @@ lean_inc(x_61); x_62 = lean_ctor_get(x_60, 1); lean_inc(x_62); lean_dec(x_60); -x_63 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9; +x_63 = l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6; lean_inc(x_2); x_64 = l_Lean_mkAppB(x_63, x_2, x_61); x_65 = l_Lean_Meta_Grind_pushEqTrue(x_2, x_64, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_62); @@ -3515,7 +3662,7 @@ return x_21; } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1() { _start: { lean_object* x_1; @@ -3523,13 +3670,13 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateNotDown), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_Grind_propagateNotUp___closed__2; x_3 = 0; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } @@ -4157,7 +4304,7 @@ return x_30; } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1() { _start: { lean_object* x_1; @@ -4165,13 +4312,13 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateEqUp), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_Grind_propagateEqUp___closed__2; x_3 = 1; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } @@ -4591,7 +4738,7 @@ return x_63; } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1() { _start: { lean_object* x_1; @@ -4599,21 +4746,29 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateEqDown), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_Grind_propagateEqUp___closed__2; x_3 = 0; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqMatchDown___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_13; +lean_object* x_14; +x_14 = l_Lean_Meta_Grind_markCaseSplitAsResolved(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -4621,21 +4776,21 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); lean_inc(x_1); -x_13 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +x_16 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; -x_17 = l_Lean_mkAppB(x_16, x_1, x_14); -x_18 = 1; -x_19 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_17, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; +x_20 = l_Lean_mkAppB(x_19, x_1, x_17); +x_21 = 0; +x_22 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_20, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4643,12 +4798,12 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_19; +return x_22; } else { -uint8_t x_20; +uint8_t x_23; +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4656,50 +4811,84 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_20 = !lean_is_exclusive(x_13); -if (x_20 == 0) +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) { -return x_13; +return x_16; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_13, 0); -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_13); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_16, 0); +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_16); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } -static lean_object* _init_l_Lean_Meta_Grind_propagateHEqDown___closed__1() { +else +{ +uint8_t x_27; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_14); +if (x_27 == 0) +{ +return x_14; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_14, 0); +x_29 = lean_ctor_get(x_14, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_14); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateEqMatchDown___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("HEq", 3, 3); +x_1 = lean_mk_string_unchecked("EqMatch", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_propagateHEqDown___closed__2() { +static lean_object* _init_l_Lean_Meta_Grind_propagateEqMatchDown___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_propagateHEqDown___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateEqMatchDown___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateEqMatchDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; @@ -4803,12 +4992,14 @@ return x_11; } else { -lean_object* x_30; uint8_t x_31; -x_30 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); -x_31 = l_Lean_Expr_isApp(x_30); -if (x_31 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l_Lean_Expr_appArg(x_27, lean_box(0)); +x_31 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); +x_32 = l_Lean_Expr_isApp(x_31); +if (x_32 == 0) { -lean_object* x_32; +lean_object* x_33; +lean_dec(x_31); lean_dec(x_30); lean_dec(x_26); lean_dec(x_9); @@ -4820,21 +5011,22 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -lean_ctor_set(x_11, 0, x_32); +x_33 = lean_box(0); +lean_ctor_set(x_11, 0, x_33); return x_11; } else { -lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_33 = l_Lean_Expr_appArg(x_30, lean_box(0)); -x_34 = l_Lean_Expr_appFnCleanup(x_30, lean_box(0)); -x_35 = l_Lean_Expr_isApp(x_34); -if (x_35 == 0) +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_35 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_36 = l_Lean_Expr_isApp(x_35); +if (x_36 == 0) { -lean_object* x_36; +lean_object* x_37; +lean_dec(x_35); lean_dec(x_34); -lean_dec(x_33); +lean_dec(x_30); lean_dec(x_26); lean_dec(x_9); lean_dec(x_8); @@ -4845,21 +5037,22 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_36 = lean_box(0); -lean_ctor_set(x_11, 0, x_36); +x_37 = lean_box(0); +lean_ctor_set(x_11, 0, x_37); return x_11; } else { -lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_37 = l_Lean_Expr_appFnCleanup(x_34, lean_box(0)); -x_38 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; -x_39 = l_Lean_Expr_isConstOf(x_37, x_38); -lean_dec(x_37); -if (x_39 == 0) +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = l_Lean_Expr_appFnCleanup(x_35, lean_box(0)); +x_39 = l_Lean_Meta_Grind_propagateEqMatchDown___closed__2; +x_40 = l_Lean_Expr_isConstOf(x_38, x_39); +lean_dec(x_38); +if (x_40 == 0) { -lean_object* x_40; -lean_dec(x_33); +lean_object* x_41; +lean_dec(x_34); +lean_dec(x_30); lean_dec(x_26); lean_dec(x_9); lean_dec(x_8); @@ -4870,16 +5063,16 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_40 = lean_box(0); -lean_ctor_set(x_11, 0, x_40); +x_41 = lean_box(0); +lean_ctor_set(x_11, 0, x_41); return x_11; } else { -lean_object* x_41; +lean_object* x_42; lean_free_object(x_11); -x_41 = l_Lean_Meta_Grind_propagateHEqDown___lambda__1(x_1, x_33, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); -return x_41; +x_42 = l_Lean_Meta_Grind_propagateEqMatchDown___lambda__1(x_1, x_34, x_30, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +return x_42; } } } @@ -4888,17 +5081,17 @@ return x_41; } else { -lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_42 = lean_ctor_get(x_11, 1); -lean_inc(x_42); +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_ctor_get(x_11, 1); +lean_inc(x_43); lean_dec(x_11); lean_inc(x_1); -x_43 = l_Lean_Expr_cleanupAnnotations(x_1); -x_44 = l_Lean_Expr_isApp(x_43); -if (x_44 == 0) +x_44 = l_Lean_Expr_cleanupAnnotations(x_1); +x_45 = l_Lean_Expr_isApp(x_44); +if (x_45 == 0) { -lean_object* x_45; lean_object* x_46; -lean_dec(x_43); +lean_object* x_46; lean_object* x_47; +lean_dec(x_44); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -4908,23 +5101,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_45 = lean_box(0); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_42); -return x_46; +x_46 = lean_box(0); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_43); +return x_47; } else { -lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_47 = l_Lean_Expr_appArg(x_43, lean_box(0)); -x_48 = l_Lean_Expr_appFnCleanup(x_43, lean_box(0)); -x_49 = l_Lean_Expr_isApp(x_48); -if (x_49 == 0) +lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_48 = l_Lean_Expr_appArg(x_44, lean_box(0)); +x_49 = l_Lean_Expr_appFnCleanup(x_44, lean_box(0)); +x_50 = l_Lean_Expr_isApp(x_49); +if (x_50 == 0) { -lean_object* x_50; lean_object* x_51; +lean_object* x_51; lean_object* x_52; +lean_dec(x_49); lean_dec(x_48); -lean_dec(x_47); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -4934,22 +5127,24 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_50 = lean_box(0); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_42); -return x_51; +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_43); +return x_52; } else { -lean_object* x_52; uint8_t x_53; -x_52 = l_Lean_Expr_appFnCleanup(x_48, lean_box(0)); -x_53 = l_Lean_Expr_isApp(x_52); -if (x_53 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = l_Lean_Expr_appArg(x_49, lean_box(0)); +x_54 = l_Lean_Expr_appFnCleanup(x_49, lean_box(0)); +x_55 = l_Lean_Expr_isApp(x_54); +if (x_55 == 0) { -lean_object* x_54; lean_object* x_55; -lean_dec(x_52); -lean_dec(x_47); +lean_object* x_56; lean_object* x_57; +lean_dec(x_54); +lean_dec(x_53); +lean_dec(x_48); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -4959,24 +5154,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_54 = lean_box(0); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_42); -return x_55; +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_43); +return x_57; } else { -lean_object* x_56; lean_object* x_57; uint8_t x_58; -x_56 = l_Lean_Expr_appArg(x_52, lean_box(0)); -x_57 = l_Lean_Expr_appFnCleanup(x_52, lean_box(0)); -x_58 = l_Lean_Expr_isApp(x_57); -if (x_58 == 0) +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = l_Lean_Expr_appArg(x_54, lean_box(0)); +x_59 = l_Lean_Expr_appFnCleanup(x_54, lean_box(0)); +x_60 = l_Lean_Expr_isApp(x_59); +if (x_60 == 0) { -lean_object* x_59; lean_object* x_60; -lean_dec(x_57); -lean_dec(x_56); -lean_dec(x_47); +lean_object* x_61; lean_object* x_62; +lean_dec(x_59); +lean_dec(x_58); +lean_dec(x_53); +lean_dec(x_48); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -4986,24 +5182,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = lean_box(0); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_42); -return x_60; +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_43); +return x_62; } else { -lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_61 = l_Lean_Expr_appFnCleanup(x_57, lean_box(0)); -x_62 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; -x_63 = l_Lean_Expr_isConstOf(x_61, x_62); -lean_dec(x_61); -if (x_63 == 0) +lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_63 = l_Lean_Expr_appFnCleanup(x_59, lean_box(0)); +x_64 = l_Lean_Meta_Grind_propagateEqMatchDown___closed__2; +x_65 = l_Lean_Expr_isConstOf(x_63, x_64); +lean_dec(x_63); +if (x_65 == 0) { -lean_object* x_64; lean_object* x_65; -lean_dec(x_56); -lean_dec(x_47); +lean_object* x_66; lean_object* x_67; +lean_dec(x_58); +lean_dec(x_53); +lean_dec(x_48); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -5013,17 +5210,17 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_64 = lean_box(0); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_42); -return x_65; +x_66 = lean_box(0); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_43); +return x_67; } else { -lean_object* x_66; -x_66 = l_Lean_Meta_Grind_propagateHEqDown___lambda__1(x_1, x_56, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); -return x_66; +lean_object* x_68; +x_68 = l_Lean_Meta_Grind_propagateEqMatchDown___lambda__1(x_1, x_58, x_53, x_48, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_43); +return x_68; } } } @@ -5034,7 +5231,7 @@ return x_66; } else { -uint8_t x_67; +uint8_t x_69; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -5044,63 +5241,2061 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_67 = !lean_is_exclusive(x_11); -if (x_67 == 0) +x_69 = !lean_is_exclusive(x_11); +if (x_69 == 0) { return x_11; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_11, 0); -x_69 = lean_ctor_get(x_11, 1); -lean_inc(x_69); -lean_inc(x_68); +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_11, 0); +x_71 = lean_ctor_get(x_11, 1); +lean_inc(x_71); +lean_inc(x_70); lean_dec(x_11); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateHEqDown), 10, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateEqMatchDown), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546_(lean_object* x_1) { -_start: +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Meta_Grind_propagateEqMatchDown___closed__2; +x_3 = 0; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1; +x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; +x_17 = l_Lean_mkAppB(x_16, x_1, x_14); +x_18 = 1; +x_19 = l_Lean_Meta_Grind_pushEqCore(x_2, x_3, x_17, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) +{ +return x_13; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_13, 0); +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_13); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateHEqDown___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("HEq", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateHEqDown___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateHEqDown___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +uint8_t x_14; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_11, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_11); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_ctor_get(x_11, 1); +x_22 = lean_ctor_get(x_11, 0); +lean_dec(x_22); +lean_inc(x_1); +x_23 = l_Lean_Expr_cleanupAnnotations(x_1); +x_24 = l_Lean_Expr_isApp(x_23); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_23); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_25 = lean_box(0); +lean_ctor_set(x_11, 0, x_25); +return x_11; +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Expr_appArg(x_23, lean_box(0)); +x_27 = l_Lean_Expr_appFnCleanup(x_23, lean_box(0)); +x_28 = l_Lean_Expr_isApp(x_27); +if (x_28 == 0) +{ +lean_object* x_29; +lean_dec(x_27); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_29 = lean_box(0); +lean_ctor_set(x_11, 0, x_29); +return x_11; +} +else +{ +lean_object* x_30; uint8_t x_31; +x_30 = l_Lean_Expr_appFnCleanup(x_27, lean_box(0)); +x_31 = l_Lean_Expr_isApp(x_30); +if (x_31 == 0) +{ +lean_object* x_32; +lean_dec(x_30); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = lean_box(0); +lean_ctor_set(x_11, 0, x_32); +return x_11; +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = l_Lean_Expr_appArg(x_30, lean_box(0)); +x_34 = l_Lean_Expr_appFnCleanup(x_30, lean_box(0)); +x_35 = l_Lean_Expr_isApp(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_box(0); +lean_ctor_set(x_11, 0, x_36); +return x_11; +} +else +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = l_Lean_Expr_appFnCleanup(x_34, lean_box(0)); +x_38 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_39 = l_Lean_Expr_isConstOf(x_37, x_38); +lean_dec(x_37); +if (x_39 == 0) +{ +lean_object* x_40; +lean_dec(x_33); +lean_dec(x_26); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_40 = lean_box(0); +lean_ctor_set(x_11, 0, x_40); +return x_11; +} +else +{ +lean_object* x_41; +lean_free_object(x_11); +x_41 = l_Lean_Meta_Grind_propagateHEqDown___lambda__1(x_1, x_33, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +return x_41; +} +} +} +} +} +} +else +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = lean_ctor_get(x_11, 1); +lean_inc(x_42); +lean_dec(x_11); +lean_inc(x_1); +x_43 = l_Lean_Expr_cleanupAnnotations(x_1); +x_44 = l_Lean_Expr_isApp(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_43); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_42); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = l_Lean_Expr_appArg(x_43, lean_box(0)); +x_48 = l_Lean_Expr_appFnCleanup(x_43, lean_box(0)); +x_49 = l_Lean_Expr_isApp(x_48); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_48); +lean_dec(x_47); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_50 = lean_box(0); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_42); +return x_51; +} +else +{ +lean_object* x_52; uint8_t x_53; +x_52 = l_Lean_Expr_appFnCleanup(x_48, lean_box(0)); +x_53 = l_Lean_Expr_isApp(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +lean_dec(x_52); +lean_dec(x_47); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_42); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_56 = l_Lean_Expr_appArg(x_52, lean_box(0)); +x_57 = l_Lean_Expr_appFnCleanup(x_52, lean_box(0)); +x_58 = l_Lean_Expr_isApp(x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_57); +lean_dec(x_56); +lean_dec(x_47); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_42); +return x_60; +} +else +{ +lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_61 = l_Lean_Expr_appFnCleanup(x_57, lean_box(0)); +x_62 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_63 = l_Lean_Expr_isConstOf(x_61, x_62); +lean_dec(x_61); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; +lean_dec(x_56); +lean_dec(x_47); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_64 = lean_box(0); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_42); +return x_65; +} +else +{ +lean_object* x_66; +x_66 = l_Lean_Meta_Grind_propagateHEqDown___lambda__1(x_1, x_56, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +return x_66; +} +} +} +} +} +} +} +} +else +{ +uint8_t x_67; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_11); +if (x_67 == 0) +{ +return x_11; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_11, 0); +x_69 = lean_ctor_get(x_11, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_11); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateHEqDown), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_3 = 0; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1; +x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_3); +lean_inc(x_2); +x_13 = l_Lean_Meta_Grind_isEqv(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +uint8_t x_16; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_23 = lean_grind_mk_heq_proof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__3; +lean_inc(x_1); +x_27 = l_Lean_mkAppB(x_26, x_1, x_24); +x_28 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_28; +} +else +{ +uint8_t x_29; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_23); +if (x_29 == 0) +{ +return x_23; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_23, 0); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_23); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +else +{ +uint8_t x_33; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_13); +if (x_33 == 0) +{ +return x_13; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_13, 0); +x_35 = lean_ctor_get(x_13, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_13); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +lean_inc(x_1); +x_11 = l_Lean_Expr_cleanupAnnotations(x_1); +x_12 = l_Lean_Expr_isApp(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_10); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Expr_appArg(x_11, lean_box(0)); +x_16 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); +x_17 = l_Lean_Expr_isApp(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_10); +return x_19; +} +else +{ +lean_object* x_20; uint8_t x_21; +x_20 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); +x_21 = l_Lean_Expr_isApp(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_10); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); +x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); +x_26 = l_Lean_Expr_isApp(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_10); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); +x_30 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_31 = l_Lean_Expr_isConstOf(x_29, x_30); +lean_dec(x_29); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_24); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_10); +return x_33; +} +else +{ +lean_object* x_34; +x_34 = l_Lean_Meta_Grind_propagateHEqUp___lambda__1(x_1, x_24, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_34; +} +} +} +} +} +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateHEqUp), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_3 = 1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1; +x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite_cond_eq_false", 17, 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite_cond_eq_true", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +lean_inc(x_4); +x_17 = l_Lean_Meta_Grind_isEqTrue(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +lean_inc(x_4); +x_21 = l_Lean_Meta_Grind_isEqFalse(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_21, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_21, 0, x_26); +return x_21; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_dec(x_21); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_31 = l_Lean_Meta_Grind_mkEqFalseProof(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_Expr_constLevels_x21(x_2); +x_35 = l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2; +x_36 = l_Lean_Expr_const___override(x_35, x_34); +lean_inc(x_7); +x_37 = l_Lean_mkApp6(x_36, x_3, x_4, x_5, x_6, x_7, x_32); +x_38 = 0; +x_39 = l_Lean_Meta_Grind_pushEqCore(x_1, x_7, x_37, x_38, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_33); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_39; +} +else +{ +uint8_t x_40; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_31); +if (x_40 == 0) +{ +return x_31; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_31, 0); +x_42 = lean_ctor_get(x_31, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_31); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +} +else +{ +uint8_t x_44; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_44 = !lean_is_exclusive(x_21); +if (x_44 == 0) +{ +return x_21; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_21, 0); +x_46 = lean_ctor_get(x_21, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_21); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_17, 1); +lean_inc(x_48); +lean_dec(x_17); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_49 = l_Lean_Meta_Grind_mkEqTrueProof(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +lean_dec(x_49); +x_52 = l_Lean_Expr_constLevels_x21(x_2); +x_53 = l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4; +x_54 = l_Lean_Expr_const___override(x_53, x_52); +lean_inc(x_6); +x_55 = l_Lean_mkApp6(x_54, x_3, x_4, x_5, x_6, x_7, x_50); +x_56 = 0; +x_57 = l_Lean_Meta_Grind_pushEqCore(x_1, x_6, x_55, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_51); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_57; +} +else +{ +uint8_t x_58; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_49); +if (x_58 == 0) +{ +return x_49; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_49, 0); +x_60 = lean_ctor_get(x_49, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_49); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +} +} +else +{ +uint8_t x_62; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_62 = !lean_is_exclusive(x_17); +if (x_62 == 0) +{ +return x_17; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_17, 0); +x_64 = lean_ctor_get(x_17, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_17); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateIte___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateIte___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +lean_inc(x_1); +x_11 = l_Lean_Expr_cleanupAnnotations(x_1); +x_12 = l_Lean_Expr_isApp(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_10); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Expr_appArg(x_11, lean_box(0)); +x_16 = l_Lean_Expr_appFnCleanup(x_11, lean_box(0)); +x_17 = l_Lean_Expr_isApp(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_10); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Expr_appArg(x_16, lean_box(0)); +x_21 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); +x_22 = l_Lean_Expr_isApp(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_10); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = l_Lean_Expr_appArg(x_21, lean_box(0)); +x_26 = l_Lean_Expr_appFnCleanup(x_21, lean_box(0)); +x_27 = l_Lean_Expr_isApp(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_20); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_10); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l_Lean_Expr_appArg(x_26, lean_box(0)); +x_31 = l_Lean_Expr_appFnCleanup(x_26, lean_box(0)); +x_32 = l_Lean_Expr_isApp(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_25); +lean_dec(x_20); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_10); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_36 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_37 = l_Lean_Meta_Grind_propagateIte___closed__2; +x_38 = l_Lean_Expr_isConstOf(x_36, x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_30); +lean_dec(x_25); +lean_dec(x_20); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_39 = lean_box(0); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_10); +return x_40; +} +else +{ +lean_object* x_41; +x_41 = l_Lean_Meta_Grind_propagateIte___lambda__1(x_1, x_36, x_35, x_30, x_25, x_20, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_36); +return x_41; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateIte___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Grind_propagateIte___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_2); +return x_17; +} +} +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateIte), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Meta_Grind_propagateIte___closed__2; +x_3 = 1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1; +x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("of_eq_false", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite_cond_eq_false'", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite_cond_eq_true'", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1; +x_2 = l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2; +x_3 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +lean_inc(x_4); +x_17 = l_Lean_Meta_Grind_isEqTrue(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +lean_inc(x_4); +x_21 = l_Lean_Meta_Grind_isEqFalse(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_21, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_21, 0, x_26); +return x_21; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_dec(x_21); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_31 = l_Lean_Meta_Grind_mkEqFalseProof(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3; +lean_inc(x_32); +lean_inc(x_4); +x_35 = l_Lean_mkAppB(x_34, x_4, x_32); +lean_inc(x_7); +x_36 = l_Lean_Expr_app___override(x_7, x_35); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_10); +x_37 = l_Lean_Meta_Grind_simp(x_36, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_33); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_41 = l_Lean_Meta_Simp_Result_getProof(x_38, x_12, x_13, x_14, x_15, x_39); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +lean_inc(x_1); +x_44 = l_Lean_Meta_Grind_getGeneration(x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_43); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_40); +x_47 = l_Lean_Meta_Grind_internalize(x_40, x_45, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_46); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_Lean_Expr_constLevels_x21(x_2); +x_50 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5; +x_51 = l_Lean_Expr_const___override(x_50, x_49); +lean_inc(x_40); +x_52 = l_Lean_mkApp8(x_51, x_3, x_4, x_5, x_6, x_7, x_40, x_32, x_42); +x_53 = 0; +x_54 = l_Lean_Meta_Grind_pushEqCore(x_1, x_40, x_52, x_53, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_48); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_54; +} +else +{ +uint8_t x_55; +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_32); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_47); +if (x_55 == 0) +{ +return x_47; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_47, 0); +x_57 = lean_ctor_get(x_47, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_47); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +else +{ +uint8_t x_59; +lean_dec(x_42); +lean_dec(x_40); +lean_dec(x_32); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_44); +if (x_59 == 0) +{ +return x_44; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_44, 0); +x_61 = lean_ctor_get(x_44, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_44); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +} +else +{ +uint8_t x_63; +lean_dec(x_40); +lean_dec(x_32); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_41); +if (x_63 == 0) +{ +return x_41; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_41, 0); +x_65 = lean_ctor_get(x_41, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_41); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +else +{ +uint8_t x_67; +lean_dec(x_32); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_37); +if (x_67 == 0) +{ +return x_37; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_37, 0); +x_69 = lean_ctor_get(x_37, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_37); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +else +{ +uint8_t x_71; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_71 = !lean_is_exclusive(x_31); +if (x_71 == 0) +{ +return x_31; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_31, 0); +x_73 = lean_ctor_get(x_31, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_31); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +} +else +{ +uint8_t x_75; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_75 = !lean_is_exclusive(x_21); +if (x_75 == 0) +{ +return x_21; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_21, 0); +x_77 = lean_ctor_get(x_21, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_21); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; +x_79 = lean_ctor_get(x_17, 1); +lean_inc(x_79); +lean_dec(x_17); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_80 = l_Lean_Meta_Grind_mkEqTrueProof(x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_79); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +lean_dec(x_80); +x_83 = l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3; +lean_inc(x_81); +lean_inc(x_4); +x_84 = l_Lean_mkAppB(x_83, x_4, x_81); +lean_inc(x_6); +x_85 = l_Lean_Expr_app___override(x_6, x_84); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_10); +x_86 = l_Lean_Meta_Grind_simp(x_85, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_82); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_90 = l_Lean_Meta_Simp_Result_getProof(x_87, x_12, x_13, x_14, x_15, x_88); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +lean_inc(x_1); +x_93 = l_Lean_Meta_Grind_getGeneration(x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_92); +if (lean_obj_tag(x_93) == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +lean_dec(x_93); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_89); +x_96 = l_Lean_Meta_Grind_internalize(x_89, x_94, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_95); +if (lean_obj_tag(x_96) == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = l_Lean_Expr_constLevels_x21(x_2); +x_99 = l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7; +x_100 = l_Lean_Expr_const___override(x_99, x_98); +lean_inc(x_89); +x_101 = l_Lean_mkApp8(x_100, x_3, x_4, x_5, x_6, x_7, x_89, x_81, x_91); +x_102 = 0; +x_103 = l_Lean_Meta_Grind_pushEqCore(x_1, x_89, x_101, x_102, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_97); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_103; +} +else +{ +uint8_t x_104; +lean_dec(x_91); +lean_dec(x_89); +lean_dec(x_81); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_104 = !lean_is_exclusive(x_96); +if (x_104 == 0) +{ +return x_96; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_96, 0); +x_106 = lean_ctor_get(x_96, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_96); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +return x_107; +} +} +} +else +{ +uint8_t x_108; +lean_dec(x_91); +lean_dec(x_89); +lean_dec(x_81); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_108 = !lean_is_exclusive(x_93); +if (x_108 == 0) +{ +return x_93; +} +else { -lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; -x_3 = 0; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1; -x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); -return x_5; +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_93, 0); +x_110 = lean_ctor_get(x_93, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_93); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -lean_inc(x_3); -lean_inc(x_2); -x_13 = l_Lean_Meta_Grind_isEqv(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +} +else { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_unbox(x_14); +uint8_t x_112; +lean_dec(x_89); +lean_dec(x_81); +lean_dec(x_15); lean_dec(x_14); -if (x_15 == 0) -{ -uint8_t x_16; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5110,58 +7305,35 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) +x_112 = !lean_is_exclusive(x_90); +if (x_112 == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = lean_box(0); -lean_ctor_set(x_13, 0, x_18); -return x_13; +return x_90; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_dec(x_13); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_90, 0); +x_114 = lean_ctor_get(x_90, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_90); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +return x_115; +} } } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); +uint8_t x_116; +lean_dec(x_81); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_23 = lean_grind_mk_heq_proof(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__3; -lean_inc(x_1); -x_27 = l_Lean_mkAppB(x_26, x_1, x_24); -x_28 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5170,11 +7342,35 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_28; +lean_dec(x_3); +lean_dec(x_1); +x_116 = !lean_is_exclusive(x_86); +if (x_116 == 0) +{ +return x_86; } else { -uint8_t x_29; +lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_117 = lean_ctor_get(x_86, 0); +x_118 = lean_ctor_get(x_86, 1); +lean_inc(x_118); +lean_inc(x_117); +lean_dec(x_86); +x_119 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_119, 0, x_117); +lean_ctor_set(x_119, 1, x_118); +return x_119; +} +} +} +else +{ +uint8_t x_120; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5183,31 +7379,36 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_29 = !lean_is_exclusive(x_23); -if (x_29 == 0) +x_120 = !lean_is_exclusive(x_80); +if (x_120 == 0) { -return x_23; +return x_80; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_23, 0); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_23); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_80, 0); +x_122 = lean_ctor_get(x_80, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_80); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; } } } } else { -uint8_t x_33; +uint8_t x_124; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -5217,30 +7418,47 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_13); -if (x_33 == 0) +x_124 = !lean_is_exclusive(x_17); +if (x_124 == 0) { -return x_13; +return x_17; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_13, 0); -x_35 = lean_ctor_get(x_13, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_13); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_17, 0); +x_126 = lean_ctor_get(x_17, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_17); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateHEqUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_propagateDIte___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_propagateDIte___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; uint8_t x_12; @@ -5294,12 +7512,14 @@ return x_19; } else { -lean_object* x_20; uint8_t x_21; -x_20 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); -x_21 = l_Lean_Expr_isApp(x_20); -if (x_21 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Expr_appArg(x_16, lean_box(0)); +x_21 = l_Lean_Expr_appFnCleanup(x_16, lean_box(0)); +x_22 = l_Lean_Expr_isApp(x_21); +if (x_22 == 0) { -lean_object* x_22; lean_object* x_23; +lean_object* x_23; lean_object* x_24; +lean_dec(x_21); lean_dec(x_20); lean_dec(x_15); lean_dec(x_9); @@ -5311,23 +7531,24 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_10); -return x_23; +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_10); +return x_24; } else { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); -x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); -x_26 = l_Lean_Expr_isApp(x_25); -if (x_26 == 0) +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = l_Lean_Expr_appArg(x_21, lean_box(0)); +x_26 = l_Lean_Expr_appFnCleanup(x_21, lean_box(0)); +x_27 = l_Lean_Expr_isApp(x_26); +if (x_27 == 0) { -lean_object* x_27; lean_object* x_28; +lean_object* x_28; lean_object* x_29; +lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); +lean_dec(x_20); lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); @@ -5338,23 +7559,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_10); -return x_28; +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_10); +return x_29; } else { -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); -x_30 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; -x_31 = l_Lean_Expr_isConstOf(x_29, x_30); -lean_dec(x_29); -if (x_31 == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l_Lean_Expr_appArg(x_26, lean_box(0)); +x_31 = l_Lean_Expr_appFnCleanup(x_26, lean_box(0)); +x_32 = l_Lean_Expr_isApp(x_31); +if (x_32 == 0) { -lean_object* x_32; lean_object* x_33; -lean_dec(x_24); +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_25); +lean_dec(x_20); lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); @@ -5365,39 +7588,81 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = lean_box(0); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_10); -return x_33; +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_10); +return x_34; } else { -lean_object* x_34; -x_34 = l_Lean_Meta_Grind_propagateHEqUp___lambda__1(x_1, x_24, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_34; +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_36 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_37 = l_Lean_Meta_Grind_propagateDIte___closed__2; +x_38 = l_Lean_Expr_isConstOf(x_36, x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_30); +lean_dec(x_25); +lean_dec(x_20); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_39 = lean_box(0); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_10); +return x_40; +} +else +{ +lean_object* x_41; +x_41 = l_Lean_Meta_Grind_propagateDIte___lambda__1(x_1, x_36, x_35, x_30, x_25, x_20, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_36); +return x_41; +} +} +} } } } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDIte___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_Lean_Meta_Grind_propagateDIte___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_2); +return x_17; } } -static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateHEqUp), 10, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_propagateDIte), 10, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Meta_Grind_propagateHEqDown___closed__2; +x_2 = l_Lean_Meta_Grind_propagateDIte___closed__2; x_3 = 1; -x_4 = l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1; +x_4 = l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1; x_5 = l___private_Lean_Meta_Tactic_Grind_PropagatorAttr_0__Lean_Meta_Grind_registerBuiltinPropagatorCore(x_2, x_3, x_4, x_1); return x_5; } @@ -5405,6 +7670,8 @@ return x_5; lean_object* initialize_Init_Grind(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Proof(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_PropagatorAttr(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Internalize(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Propagate(uint8_t builtin, lean_object* w) { lean_object * res; @@ -5419,6 +7686,12 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_PropagatorAttr(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Internalize(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__1); l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateAndUp___lambda__1___closed__2(); @@ -5537,13 +7810,19 @@ l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__5 = _init_l_Lean_Meta_Gri lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__5); l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6(); lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__6); +l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__7); +l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__8); +l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9 = _init_l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___lambda__1___closed__9); l_Lean_Meta_Grind_propagateNotUp___closed__1 = _init_l_Lean_Meta_Grind_propagateNotUp___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___closed__1); l_Lean_Meta_Grind_propagateNotUp___closed__2 = _init_l_Lean_Meta_Grind_propagateNotUp___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_propagateNotUp___closed__2); -l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1426_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateNotUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1486_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__1(); @@ -5558,15 +7837,9 @@ l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__5 = _init_l_Lean_Meta_G lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__5); l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6(); lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__6); -l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__7); -l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__8); -l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9 = _init_l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9(); -lean_mark_persistent(l_Lean_Meta_Grind_propagateNotDown___lambda__1___closed__9); -l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1706_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateNotDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_1766_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateEqUp___lambda__1___closed__1(); @@ -5591,9 +7864,9 @@ l_Lean_Meta_Grind_propagateEqUp___closed__1 = _init_l_Lean_Meta_Grind_propagateE lean_mark_persistent(l_Lean_Meta_Grind_propagateEqUp___closed__1); l_Lean_Meta_Grind_propagateEqUp___closed__2 = _init_l_Lean_Meta_Grind_propagateEqUp___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_propagateEqUp___closed__2); -l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2058_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2118_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__1(); @@ -5602,23 +7875,72 @@ l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__2 = _init_l_Lean_Meta_Gr lean_mark_persistent(l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__2); l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3(); lean_mark_persistent(l_Lean_Meta_Grind_propagateEqDown___lambda__1___closed__3); -l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2287_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2347_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_Meta_Grind_propagateEqMatchDown___closed__1 = _init_l_Lean_Meta_Grind_propagateEqMatchDown___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateEqMatchDown___closed__1); +l_Lean_Meta_Grind_propagateEqMatchDown___closed__2 = _init_l_Lean_Meta_Grind_propagateEqMatchDown___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateEqMatchDown___closed__2); +l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateEqMatchDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2629_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Grind_propagateHEqDown___closed__1 = _init_l_Lean_Meta_Grind_propagateHEqDown___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_propagateHEqDown___closed__1); l_Lean_Meta_Grind_propagateHEqDown___closed__2 = _init_l_Lean_Meta_Grind_propagateHEqDown___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_propagateHEqDown___closed__2); -l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2546_(lean_io_mk_world()); +l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateHEqDown_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2888_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3148_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +}l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___lambda__1___closed__3); +l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___lambda__1___closed__4); +l_Lean_Meta_Grind_propagateIte___closed__1 = _init_l_Lean_Meta_Grind_propagateIte___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___closed__1); +l_Lean_Meta_Grind_propagateIte___closed__2 = _init_l_Lean_Meta_Grind_propagateIte___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateIte___closed__2); +l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_3530_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806____closed__1); -if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateHEqUp_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_2806_(lean_io_mk_world()); +}l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__1); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__2); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__3); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__4); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__5); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__6); +l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7 = _init_l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___lambda__1___closed__7); +l_Lean_Meta_Grind_propagateDIte___closed__1 = _init_l_Lean_Meta_Grind_propagateDIte___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___closed__1); +l_Lean_Meta_Grind_propagateDIte___closed__2 = _init_l_Lean_Meta_Grind_propagateDIte___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_propagateDIte___closed__2); +l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1 = _init_l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078____closed__1); +if (builtin) {res = l___regBuiltin_Lean_Meta_Grind_propagateDIte_declare__1____x40_Lean_Meta_Tactic_Grind_Propagate___hyg_4078_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c index 3e5d9b326863..ba110f27abd9 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Simp.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Grind.Simp -// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Simp.Main Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.MarkNestedProofs +// Imports: Init.Grind.Lemmas Lean.Meta.Tactic.Assert Lean.Meta.Tactic.Simp.Main Lean.Meta.Tactic.Grind.Util Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.DoNotSimp Lean.Meta.Tactic.Grind.MarkNestedProofs #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -21,6 +21,7 @@ lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__2___boxed(lean_obj lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simp___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_normalizeLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_simp___closed__12; lean_object* l_Lean_stringToMessageData(lean_object*); @@ -42,12 +43,16 @@ lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_obje lean_object* l_Lean_Meta_Grind_markNestedProofsImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_simp___closed__14; +lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Meta_Grind_simp___closed__5; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_shareCommon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_simp___closed__3; +static lean_object* l_Lean_Meta_Grind_simp___closed__15; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Grind_simp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -61,14 +66,13 @@ lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_obj static double l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3___closed__1; static lean_object* l_Lean_Meta_Grind_simp___closed__8; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_simp___closed__10; static lean_object* l_Lean_Meta_Grind_simp___closed__6; static lean_object* l_Lean_Meta_Grind_simp___closed__13; lean_object* l_Lean_Meta_Grind_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Grind_simp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_simpCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -139,13 +143,15 @@ return x_31; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; x_32 = lean_ctor_get(x_23, 0); x_33 = lean_ctor_get(x_23, 1); x_34 = lean_ctor_get(x_23, 2); x_35 = lean_ctor_get(x_23, 3); x_36 = lean_ctor_get(x_23, 5); x_37 = lean_ctor_get(x_23, 6); +x_38 = lean_ctor_get(x_23, 7); +lean_inc(x_38); lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); @@ -153,55 +159,56 @@ lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_dec(x_23); -x_38 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_38, 0, x_32); -lean_ctor_set(x_38, 1, x_33); -lean_ctor_set(x_38, 2, x_34); -lean_ctor_set(x_38, 3, x_35); -lean_ctor_set(x_38, 4, x_21); -lean_ctor_set(x_38, 5, x_36); -lean_ctor_set(x_38, 6, x_37); -x_39 = lean_st_ref_set(x_4, x_38, x_24); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; +x_39 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_39, 0, x_32); +lean_ctor_set(x_39, 1, x_33); +lean_ctor_set(x_39, 2, x_34); +lean_ctor_set(x_39, 3, x_35); +lean_ctor_set(x_39, 4, x_21); +lean_ctor_set(x_39, 5, x_36); +lean_ctor_set(x_39, 6, x_37); +lean_ctor_set(x_39, 7, x_38); +x_40 = lean_st_ref_set(x_4, x_39, x_24); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + x_42 = x_40; } else { - lean_dec_ref(x_39); - x_41 = lean_box(0); + lean_dec_ref(x_40); + x_42 = lean_box(0); } -if (lean_is_scalar(x_41)) { - x_42 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_42)) { + x_43 = lean_alloc_ctor(0, 2, 0); } else { - x_42 = x_41; + x_43 = x_42; } -lean_ctor_set(x_42, 0, x_20); -lean_ctor_set(x_42, 1, x_40); -return x_42; +lean_ctor_set(x_43, 0, x_20); +lean_ctor_set(x_43, 1, x_41); +return x_43; } } else { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_17); -if (x_43 == 0) +uint8_t x_44; +x_44 = !lean_is_exclusive(x_17); +if (x_44 == 0) { return x_17; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_17, 0); -x_45 = lean_ctor_get(x_17, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_17, 0); +x_46 = lean_ctor_get(x_17, 1); +lean_inc(x_46); lean_inc(x_45); -lean_inc(x_44); lean_dec(x_17); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } @@ -744,7 +751,7 @@ static lean_object* _init_l_Lean_Meta_Grind_simp___closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__3___boxed), 6, 0); return x_1; } } @@ -752,7 +759,7 @@ static lean_object* _init_l_Lean_Meta_Grind_simp___closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__1), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2), 6, 0); return x_1; } } @@ -768,7 +775,7 @@ static lean_object* _init_l_Lean_Meta_Grind_simp___closed__8() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("grind", 5, 5); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__3), 4, 0); return x_1; } } @@ -776,21 +783,37 @@ static lean_object* _init_l_Lean_Meta_Grind_simp___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("simp", 4, 4); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_eraseDoNotSimp___lambda__4___boxed), 4, 0); return x_1; } } static lean_object* _init_l_Lean_Meta_Grind_simp___closed__10() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("simp", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__12() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Grind_simp___closed__8; -x_2 = l_Lean_Meta_Grind_simp___closed__9; +x_1 = l_Lean_Meta_Grind_simp___closed__10; +x_2 = l_Lean_Meta_Grind_simp___closed__11; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Grind_simp___closed__11() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -799,7 +822,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Grind_simp___closed__12() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__14() { _start: { lean_object* x_1; @@ -807,11 +830,11 @@ x_1 = lean_mk_string_unchecked("\n===>\n", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_Grind_simp___closed__13() { +static lean_object* _init_l_Lean_Meta_Grind_simp___closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_simp___closed__12; +x_1 = l_Lean_Meta_Grind_simp___closed__14; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -924,224 +947,237 @@ lean_inc(x_7); x_43 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_40, x_42, x_32, x_7, x_8, x_41); if (lean_obj_tag(x_43) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; x_44 = lean_ctor_get(x_43, 0); lean_inc(x_44); x_45 = lean_ctor_get(x_43, 1); lean_inc(x_45); lean_dec(x_43); +x_46 = l_Lean_Meta_Grind_simp___closed__8; +x_47 = l_Lean_Meta_Grind_simp___closed__9; +lean_inc(x_8); +lean_inc(x_7); +x_48 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_44, x_46, x_47, x_7, x_8, x_45); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_46 = l_Lean_Meta_Grind_canon(x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_45); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); -lean_dec(x_46); -x_49 = l_Lean_Meta_Grind_shareCommon(x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_49, 1); -x_53 = l_Lean_Meta_Grind_simp___closed__10; -x_54 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_53, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_52); -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_unbox(x_55); -lean_dec(x_55); -if (x_56 == 0) +x_51 = l_Lean_Meta_Grind_canon(x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_50); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_Meta_Grind_shareCommon(x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_53); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_56 = lean_ctor_get(x_54, 0); +x_57 = lean_ctor_get(x_54, 1); +x_58 = l_Lean_Meta_Grind_simp___closed__12; +x_59 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_57); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_unbox(x_60); +lean_dec(x_60); +if (x_61 == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -lean_free_object(x_49); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_free_object(x_54); lean_free_object(x_10); lean_dec(x_12); -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -lean_dec(x_54); -x_58 = lean_box(0); -x_59 = l_Lean_Meta_Grind_simp___lambda__1(x_51, x_18, x_19, x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_57); +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +lean_dec(x_59); +x_63 = lean_box(0); +x_64 = l_Lean_Meta_Grind_simp___lambda__1(x_56, x_18, x_19, x_63, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_62); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -return x_59; +return x_64; } else { -uint8_t x_60; -x_60 = !lean_is_exclusive(x_54); -if (x_60 == 0) -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_61 = lean_ctor_get(x_54, 1); -x_62 = lean_ctor_get(x_54, 0); -lean_dec(x_62); -x_63 = l_Lean_MessageData_ofExpr(x_12); -x_64 = l_Lean_Meta_Grind_simp___closed__11; +uint8_t x_65; +x_65 = !lean_is_exclusive(x_59); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_66 = lean_ctor_get(x_59, 1); +x_67 = lean_ctor_get(x_59, 0); +lean_dec(x_67); +x_68 = l_Lean_MessageData_ofExpr(x_12); +x_69 = l_Lean_Meta_Grind_simp___closed__13; +lean_ctor_set_tag(x_59, 7); +lean_ctor_set(x_59, 1, x_68); +lean_ctor_set(x_59, 0, x_69); +x_70 = l_Lean_Meta_Grind_simp___closed__15; lean_ctor_set_tag(x_54, 7); -lean_ctor_set(x_54, 1, x_63); -lean_ctor_set(x_54, 0, x_64); -x_65 = l_Lean_Meta_Grind_simp___closed__13; -lean_ctor_set_tag(x_49, 7); -lean_ctor_set(x_49, 1, x_65); -lean_ctor_set(x_49, 0, x_54); -lean_inc(x_51); -x_66 = l_Lean_MessageData_ofExpr(x_51); +lean_ctor_set(x_54, 1, x_70); +lean_ctor_set(x_54, 0, x_59); +lean_inc(x_56); +x_71 = l_Lean_MessageData_ofExpr(x_56); lean_ctor_set_tag(x_10, 7); -lean_ctor_set(x_10, 1, x_66); -lean_ctor_set(x_10, 0, x_49); -x_67 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_67, 0, x_10); -lean_ctor_set(x_67, 1, x_64); -x_68 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_53, x_67, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_61); -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Lean_Meta_Grind_simp___lambda__1(x_51, x_18, x_19, x_69, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_70); +lean_ctor_set(x_10, 1, x_71); +lean_ctor_set(x_10, 0, x_54); +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_10); +lean_ctor_set(x_72, 1, x_69); +x_73 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_58, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_66); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_Grind_simp___lambda__1(x_56, x_18, x_19, x_74, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_75); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_69); -return x_71; +lean_dec(x_74); +return x_76; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_72 = lean_ctor_get(x_54, 1); -lean_inc(x_72); -lean_dec(x_54); -x_73 = l_Lean_MessageData_ofExpr(x_12); -x_74 = l_Lean_Meta_Grind_simp___closed__11; -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_73); -x_76 = l_Lean_Meta_Grind_simp___closed__13; -lean_ctor_set_tag(x_49, 7); -lean_ctor_set(x_49, 1, x_76); -lean_ctor_set(x_49, 0, x_75); -lean_inc(x_51); -x_77 = l_Lean_MessageData_ofExpr(x_51); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_77 = lean_ctor_get(x_59, 1); +lean_inc(x_77); +lean_dec(x_59); +x_78 = l_Lean_MessageData_ofExpr(x_12); +x_79 = l_Lean_Meta_Grind_simp___closed__13; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_78); +x_81 = l_Lean_Meta_Grind_simp___closed__15; +lean_ctor_set_tag(x_54, 7); +lean_ctor_set(x_54, 1, x_81); +lean_ctor_set(x_54, 0, x_80); +lean_inc(x_56); +x_82 = l_Lean_MessageData_ofExpr(x_56); lean_ctor_set_tag(x_10, 7); -lean_ctor_set(x_10, 1, x_77); -lean_ctor_set(x_10, 0, x_49); -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_10); -lean_ctor_set(x_78, 1, x_74); -x_79 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_53, x_78, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_72); -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -lean_dec(x_79); -x_82 = l_Lean_Meta_Grind_simp___lambda__1(x_51, x_18, x_19, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_81); +lean_ctor_set(x_10, 1, x_82); +lean_ctor_set(x_10, 0, x_54); +x_83 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_83, 0, x_10); +lean_ctor_set(x_83, 1, x_79); +x_84 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_58, x_83, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_77); +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = l_Lean_Meta_Grind_simp___lambda__1(x_56, x_18, x_19, x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_86); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_80); -return x_82; +lean_dec(x_85); +return x_87; } } } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_83 = lean_ctor_get(x_49, 0); -x_84 = lean_ctor_get(x_49, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_49); -x_85 = l_Lean_Meta_Grind_simp___closed__10; -x_86 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_84); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_unbox(x_87); -lean_dec(x_87); -if (x_88 == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_88 = lean_ctor_get(x_54, 0); +x_89 = lean_ctor_get(x_54, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_54); +x_90 = l_Lean_Meta_Grind_simp___closed__12; +x_91 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_90, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_89); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_unbox(x_92); +lean_dec(x_92); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_free_object(x_10); lean_dec(x_12); -x_89 = lean_ctor_get(x_86, 1); -lean_inc(x_89); -lean_dec(x_86); -x_90 = lean_box(0); -x_91 = l_Lean_Meta_Grind_simp___lambda__1(x_83, x_18, x_19, x_90, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_89); +x_94 = lean_ctor_get(x_91, 1); +lean_inc(x_94); +lean_dec(x_91); +x_95 = lean_box(0); +x_96 = l_Lean_Meta_Grind_simp___lambda__1(x_88, x_18, x_19, x_95, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_94); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -return x_91; +return x_96; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_92 = lean_ctor_get(x_86, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_86)) { - lean_ctor_release(x_86, 0); - lean_ctor_release(x_86, 1); - x_93 = x_86; +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_97 = lean_ctor_get(x_91, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_98 = x_91; } else { - lean_dec_ref(x_86); - x_93 = lean_box(0); + lean_dec_ref(x_91); + x_98 = lean_box(0); } -x_94 = l_Lean_MessageData_ofExpr(x_12); -x_95 = l_Lean_Meta_Grind_simp___closed__11; -if (lean_is_scalar(x_93)) { - x_96 = lean_alloc_ctor(7, 2, 0); +x_99 = l_Lean_MessageData_ofExpr(x_12); +x_100 = l_Lean_Meta_Grind_simp___closed__13; +if (lean_is_scalar(x_98)) { + x_101 = lean_alloc_ctor(7, 2, 0); } else { - x_96 = x_93; - lean_ctor_set_tag(x_96, 7); -} -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_94); -x_97 = l_Lean_Meta_Grind_simp___closed__13; -x_98 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -lean_inc(x_83); -x_99 = l_Lean_MessageData_ofExpr(x_83); + x_101 = x_98; + lean_ctor_set_tag(x_101, 7); +} +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_99); +x_102 = l_Lean_Meta_Grind_simp___closed__15; +x_103 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +lean_inc(x_88); +x_104 = l_Lean_MessageData_ofExpr(x_88); lean_ctor_set_tag(x_10, 7); -lean_ctor_set(x_10, 1, x_99); -lean_ctor_set(x_10, 0, x_98); -x_100 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_100, 0, x_10); -lean_ctor_set(x_100, 1, x_95); -x_101 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_85, x_100, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_92); -x_102 = lean_ctor_get(x_101, 0); -lean_inc(x_102); -x_103 = lean_ctor_get(x_101, 1); -lean_inc(x_103); -lean_dec(x_101); -x_104 = l_Lean_Meta_Grind_simp___lambda__1(x_83, x_18, x_19, x_102, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_103); +lean_ctor_set(x_10, 1, x_104); +lean_ctor_set(x_10, 0, x_103); +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_10); +lean_ctor_set(x_105, 1, x_100); +x_106 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_90, x_105, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_97); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = l_Lean_Meta_Grind_simp___lambda__1(x_88, x_18, x_19, x_107, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_108); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_102); -return x_104; +lean_dec(x_107); +return x_109; } } } else { -uint8_t x_105; +uint8_t x_110; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1150,29 +1186,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_105 = !lean_is_exclusive(x_46); -if (x_105 == 0) +x_110 = !lean_is_exclusive(x_51); +if (x_110 == 0) { -return x_46; +return x_51; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_46, 0); -x_107 = lean_ctor_get(x_46, 1); -lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_46); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_51, 0); +x_112 = lean_ctor_get(x_51, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_51); +x_113 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +return x_113; } } } else { -uint8_t x_109; +uint8_t x_114; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1181,29 +1217,60 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_109 = !lean_is_exclusive(x_43); -if (x_109 == 0) +x_114 = !lean_is_exclusive(x_48); +if (x_114 == 0) +{ +return x_48; +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_48, 0); +x_116 = lean_ctor_get(x_48, 1); +lean_inc(x_116); +lean_inc(x_115); +lean_dec(x_48); +x_117 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +return x_117; +} +} +} +else +{ +uint8_t x_118; +lean_dec(x_18); +lean_free_object(x_10); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_118 = !lean_is_exclusive(x_43); +if (x_118 == 0) { return x_43; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_43, 0); -x_111 = lean_ctor_get(x_43, 1); -lean_inc(x_111); -lean_inc(x_110); +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_43, 0); +x_120 = lean_ctor_get(x_43, 1); +lean_inc(x_120); +lean_inc(x_119); lean_dec(x_43); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +x_121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +return x_121; } } } else { -uint8_t x_113; +uint8_t x_122; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1212,29 +1279,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_113 = !lean_is_exclusive(x_39); -if (x_113 == 0) +x_122 = !lean_is_exclusive(x_39); +if (x_122 == 0) { return x_39; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_39, 0); -x_115 = lean_ctor_get(x_39, 1); -lean_inc(x_115); -lean_inc(x_114); +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_39, 0); +x_124 = lean_ctor_get(x_39, 1); +lean_inc(x_124); +lean_inc(x_123); lean_dec(x_39); -x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -return x_116; +x_125 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; } } } else { -uint8_t x_117; +uint8_t x_126; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1243,29 +1310,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_117 = !lean_is_exclusive(x_33); -if (x_117 == 0) +x_126 = !lean_is_exclusive(x_33); +if (x_126 == 0) { return x_33; } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_33, 0); -x_119 = lean_ctor_get(x_33, 1); -lean_inc(x_119); -lean_inc(x_118); +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_33, 0); +x_128 = lean_ctor_get(x_33, 1); +lean_inc(x_128); +lean_inc(x_127); lean_dec(x_33); -x_120 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set(x_120, 1, x_119); -return x_120; +x_129 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +return x_129; } } } else { -uint8_t x_121; +uint8_t x_130; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1274,29 +1341,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_121 = !lean_is_exclusive(x_28); -if (x_121 == 0) +x_130 = !lean_is_exclusive(x_28); +if (x_130 == 0) { return x_28; } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_28, 0); -x_123 = lean_ctor_get(x_28, 1); -lean_inc(x_123); -lean_inc(x_122); +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_28, 0); +x_132 = lean_ctor_get(x_28, 1); +lean_inc(x_132); +lean_inc(x_131); lean_dec(x_28); -x_124 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); -return x_124; +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +return x_133; } } } else { -uint8_t x_125; +uint8_t x_134; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1305,29 +1372,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_125 = !lean_is_exclusive(x_23); -if (x_125 == 0) +x_134 = !lean_is_exclusive(x_23); +if (x_134 == 0) { return x_23; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_23, 0); -x_127 = lean_ctor_get(x_23, 1); -lean_inc(x_127); -lean_inc(x_126); +lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_135 = lean_ctor_get(x_23, 0); +x_136 = lean_ctor_get(x_23, 1); +lean_inc(x_136); +lean_inc(x_135); lean_dec(x_23); -x_128 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -return x_128; +x_137 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +return x_137; } } } else { -uint8_t x_129; +uint8_t x_138; lean_dec(x_18); lean_free_object(x_10); lean_dec(x_12); @@ -1336,29 +1403,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_129 = !lean_is_exclusive(x_20); -if (x_129 == 0) +x_138 = !lean_is_exclusive(x_20); +if (x_138 == 0) { return x_20; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_20, 0); -x_131 = lean_ctor_get(x_20, 1); -lean_inc(x_131); -lean_inc(x_130); +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_20, 0); +x_140 = lean_ctor_get(x_20, 1); +lean_inc(x_140); +lean_inc(x_139); lean_dec(x_20); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -return x_132; +x_141 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +return x_141; } } } else { -uint8_t x_133; +uint8_t x_142; lean_free_object(x_10); lean_dec(x_12); lean_dec(x_8); @@ -1366,495 +1433,540 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_133 = !lean_is_exclusive(x_14); -if (x_133 == 0) +x_142 = !lean_is_exclusive(x_14); +if (x_142 == 0) { return x_14; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = lean_ctor_get(x_14, 0); -x_135 = lean_ctor_get(x_14, 1); -lean_inc(x_135); -lean_inc(x_134); +lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_143 = lean_ctor_get(x_14, 0); +x_144 = lean_ctor_get(x_14, 1); +lean_inc(x_144); +lean_inc(x_143); lean_dec(x_14); -x_136 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_135); -return x_136; +x_145 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_145, 0, x_143); +lean_ctor_set(x_145, 1, x_144); +return x_145; } } } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_10, 0); -x_138 = lean_ctor_get(x_10, 1); -lean_inc(x_138); -lean_inc(x_137); +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_10, 0); +x_147 = lean_ctor_get(x_10, 1); +lean_inc(x_147); +lean_inc(x_146); lean_dec(x_10); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); -lean_inc(x_137); -x_139 = l_Lean_Meta_Grind_simpCore(x_137, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_138); -if (lean_obj_tag(x_139) == 0) -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; lean_object* x_145; -x_140 = lean_ctor_get(x_139, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_139, 1); -lean_inc(x_141); -lean_dec(x_139); -x_142 = lean_ctor_get(x_140, 0); -lean_inc(x_142); -x_143 = lean_ctor_get(x_140, 1); -lean_inc(x_143); -x_144 = lean_ctor_get_uint8(x_140, sizeof(void*)*2); -lean_dec(x_140); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_145 = l_Lean_Meta_Grind_abstractNestedProofs(x_142, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_141); -if (lean_obj_tag(x_145) == 0) -{ -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_145, 0); lean_inc(x_146); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -lean_dec(x_145); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_148 = l_Lean_Meta_Grind_markNestedProofsImpl(x_146, x_5, x_6, x_7, x_8, x_147); +x_148 = l_Lean_Meta_Grind_simpCore(x_146, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_147); if (lean_obj_tag(x_148) == 0) { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_object* x_154; x_149 = lean_ctor_get(x_148, 0); lean_inc(x_149); x_150 = lean_ctor_get(x_148, 1); lean_inc(x_150); lean_dec(x_148); -x_151 = l_Lean_Meta_Grind_simp___closed__1; -x_152 = l_Lean_Meta_Grind_simp___closed__2; +x_151 = lean_ctor_get(x_149, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_149, 1); +lean_inc(x_152); +x_153 = lean_ctor_get_uint8(x_149, sizeof(void*)*2); +lean_dec(x_149); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_153 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_149, x_151, x_152, x_5, x_6, x_7, x_8, x_150); -if (lean_obj_tag(x_153) == 0) +x_154 = l_Lean_Meta_Grind_abstractNestedProofs(x_151, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_150); +if (lean_obj_tag(x_154) == 0) { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_153, 1); +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_154, 0); lean_inc(x_155); -lean_dec(x_153); -x_156 = l_Lean_Meta_Grind_simp___closed__3; -x_157 = l_Lean_Meta_Grind_simp___closed__4; +x_156 = lean_ctor_get(x_154, 1); +lean_inc(x_156); +lean_dec(x_154); lean_inc(x_8); lean_inc(x_7); -x_158 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_154, x_156, x_157, x_7, x_8, x_155); -if (lean_obj_tag(x_158) == 0) +lean_inc(x_6); +lean_inc(x_5); +x_157 = l_Lean_Meta_Grind_markNestedProofsImpl(x_155, x_5, x_6, x_7, x_8, x_156); +if (lean_obj_tag(x_157) == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; -x_159 = lean_ctor_get(x_158, 0); +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_158 = lean_ctor_get(x_157, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_157, 1); lean_inc(x_159); -x_160 = lean_ctor_get(x_158, 1); -lean_inc(x_160); -lean_dec(x_158); -x_161 = l_Lean_Meta_Grind_simp___closed__5; -x_162 = l_Lean_Meta_Grind_simp___closed__6; -x_163 = 0; +lean_dec(x_157); +x_160 = l_Lean_Meta_Grind_simp___closed__1; +x_161 = l_Lean_Meta_Grind_simp___closed__2; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_164 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_159, x_161, x_162, x_163, x_163, x_5, x_6, x_7, x_8, x_160); -if (lean_obj_tag(x_164) == 0) -{ -lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_164, 1); -lean_inc(x_166); -lean_dec(x_164); -x_167 = l_Lean_Meta_Grind_simp___closed__7; +x_162 = l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1(x_158, x_160, x_161, x_5, x_6, x_7, x_8, x_159); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = l_Lean_Meta_Grind_simp___closed__3; +x_166 = l_Lean_Meta_Grind_simp___closed__4; lean_inc(x_8); lean_inc(x_7); -x_168 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_165, x_167, x_157, x_7, x_8, x_166); -if (lean_obj_tag(x_168) == 0) +x_167 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_163, x_165, x_166, x_7, x_8, x_164); +if (lean_obj_tag(x_167) == 0) { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_168, 0); +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; lean_object* x_173; +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_167, 1); lean_inc(x_169); -x_170 = lean_ctor_get(x_168, 1); -lean_inc(x_170); -lean_dec(x_168); +lean_dec(x_167); +x_170 = l_Lean_Meta_Grind_simp___closed__5; +x_171 = l_Lean_Meta_Grind_simp___closed__6; +x_172 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_171 = l_Lean_Meta_Grind_canon(x_169, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_170); -if (lean_obj_tag(x_171) == 0) -{ -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; -x_172 = lean_ctor_get(x_171, 0); -lean_inc(x_172); -x_173 = lean_ctor_get(x_171, 1); -lean_inc(x_173); -lean_dec(x_171); -x_174 = l_Lean_Meta_Grind_shareCommon(x_172, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_173); -x_175 = lean_ctor_get(x_174, 0); +x_173 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_168, x_170, x_171, x_172, x_172, x_5, x_6, x_7, x_8, x_169); +if (lean_obj_tag(x_173) == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_173, 1); lean_inc(x_175); -x_176 = lean_ctor_get(x_174, 1); -lean_inc(x_176); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_177 = x_174; -} else { - lean_dec_ref(x_174); - x_177 = lean_box(0); -} -x_178 = l_Lean_Meta_Grind_simp___closed__10; -x_179 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_178, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_176); -x_180 = lean_ctor_get(x_179, 0); -lean_inc(x_180); -x_181 = lean_unbox(x_180); -lean_dec(x_180); -if (x_181 == 0) -{ -lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_173); +x_176 = l_Lean_Meta_Grind_simp___closed__7; +lean_inc(x_8); +lean_inc(x_7); +x_177 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_174, x_176, x_166, x_7, x_8, x_175); +if (lean_obj_tag(x_177) == 0) +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_178 = lean_ctor_get(x_177, 0); +lean_inc(x_178); +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); lean_dec(x_177); -lean_dec(x_137); -x_182 = lean_ctor_get(x_179, 1); -lean_inc(x_182); -lean_dec(x_179); -x_183 = lean_box(0); -x_184 = l_Lean_Meta_Grind_simp___lambda__1(x_175, x_143, x_144, x_183, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_182); +x_180 = l_Lean_Meta_Grind_simp___closed__8; +x_181 = l_Lean_Meta_Grind_simp___closed__9; +lean_inc(x_8); +lean_inc(x_7); +x_182 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_178, x_180, x_181, x_7, x_8, x_179); +if (lean_obj_tag(x_182) == 0) +{ +lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_183 = lean_ctor_get(x_182, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_182, 1); +lean_inc(x_184); +lean_dec(x_182); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_185 = l_Lean_Meta_Grind_canon(x_183, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_184); +if (lean_obj_tag(x_185) == 0) +{ +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; uint8_t x_195; +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_185, 1); +lean_inc(x_187); +lean_dec(x_185); +x_188 = l_Lean_Meta_Grind_shareCommon(x_186, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_187); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +if (lean_is_exclusive(x_188)) { + lean_ctor_release(x_188, 0); + lean_ctor_release(x_188, 1); + x_191 = x_188; +} else { + lean_dec_ref(x_188); + x_191 = lean_box(0); +} +x_192 = l_Lean_Meta_Grind_simp___closed__12; +x_193 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_simp___spec__2(x_192, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_190); +x_194 = lean_ctor_get(x_193, 0); +lean_inc(x_194); +x_195 = lean_unbox(x_194); +lean_dec(x_194); +if (x_195 == 0) +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_191); +lean_dec(x_146); +x_196 = lean_ctor_get(x_193, 1); +lean_inc(x_196); +lean_dec(x_193); +x_197 = lean_box(0); +x_198 = l_Lean_Meta_Grind_simp___lambda__1(x_189, x_152, x_153, x_197, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_196); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -return x_184; +return x_198; } else { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_185 = lean_ctor_get(x_179, 1); -lean_inc(x_185); -if (lean_is_exclusive(x_179)) { - lean_ctor_release(x_179, 0); - lean_ctor_release(x_179, 1); - x_186 = x_179; +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_199 = lean_ctor_get(x_193, 1); +lean_inc(x_199); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + x_200 = x_193; } else { - lean_dec_ref(x_179); - x_186 = lean_box(0); + lean_dec_ref(x_193); + x_200 = lean_box(0); } -x_187 = l_Lean_MessageData_ofExpr(x_137); -x_188 = l_Lean_Meta_Grind_simp___closed__11; -if (lean_is_scalar(x_186)) { - x_189 = lean_alloc_ctor(7, 2, 0); +x_201 = l_Lean_MessageData_ofExpr(x_146); +x_202 = l_Lean_Meta_Grind_simp___closed__13; +if (lean_is_scalar(x_200)) { + x_203 = lean_alloc_ctor(7, 2, 0); } else { - x_189 = x_186; - lean_ctor_set_tag(x_189, 7); -} -lean_ctor_set(x_189, 0, x_188); -lean_ctor_set(x_189, 1, x_187); -x_190 = l_Lean_Meta_Grind_simp___closed__13; -if (lean_is_scalar(x_177)) { - x_191 = lean_alloc_ctor(7, 2, 0); + x_203 = x_200; + lean_ctor_set_tag(x_203, 7); +} +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_201); +x_204 = l_Lean_Meta_Grind_simp___closed__15; +if (lean_is_scalar(x_191)) { + x_205 = lean_alloc_ctor(7, 2, 0); } else { - x_191 = x_177; - lean_ctor_set_tag(x_191, 7); + x_205 = x_191; + lean_ctor_set_tag(x_205, 7); +} +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +lean_inc(x_189); +x_206 = l_Lean_MessageData_ofExpr(x_189); +x_207 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_207, 0, x_205); +lean_ctor_set(x_207, 1, x_206); +x_208 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_208, 0, x_207); +lean_ctor_set(x_208, 1, x_202); +x_209 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_192, x_208, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_199); +x_210 = lean_ctor_get(x_209, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_209, 1); +lean_inc(x_211); +lean_dec(x_209); +x_212 = l_Lean_Meta_Grind_simp___lambda__1(x_189, x_152, x_153, x_210, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_211); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_210); +return x_212; } -lean_ctor_set(x_191, 0, x_189); -lean_ctor_set(x_191, 1, x_190); -lean_inc(x_175); -x_192 = l_Lean_MessageData_ofExpr(x_175); -x_193 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_193, 0, x_191); -lean_ctor_set(x_193, 1, x_192); -x_194 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_194, 0, x_193); -lean_ctor_set(x_194, 1, x_188); -x_195 = l_Lean_addTrace___at_Lean_Meta_Grind_simp___spec__3(x_178, x_194, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_185); -x_196 = lean_ctor_get(x_195, 0); -lean_inc(x_196); -x_197 = lean_ctor_get(x_195, 1); -lean_inc(x_197); -lean_dec(x_195); -x_198 = l_Lean_Meta_Grind_simp___lambda__1(x_175, x_143, x_144, x_196, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_197); +} +else +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_196); -return x_198; +x_213 = lean_ctor_get(x_185, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_185, 1); +lean_inc(x_214); +if (lean_is_exclusive(x_185)) { + lean_ctor_release(x_185, 0); + lean_ctor_release(x_185, 1); + x_215 = x_185; +} else { + lean_dec_ref(x_185); + x_215 = lean_box(0); +} +if (lean_is_scalar(x_215)) { + x_216 = lean_alloc_ctor(1, 2, 0); +} else { + x_216 = x_215; +} +lean_ctor_set(x_216, 0, x_213); +lean_ctor_set(x_216, 1, x_214); +return x_216; } } else { -lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_199 = lean_ctor_get(x_171, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_171, 1); -lean_inc(x_200); -if (lean_is_exclusive(x_171)) { - lean_ctor_release(x_171, 0); - lean_ctor_release(x_171, 1); - x_201 = x_171; +x_217 = lean_ctor_get(x_182, 0); +lean_inc(x_217); +x_218 = lean_ctor_get(x_182, 1); +lean_inc(x_218); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + x_219 = x_182; } else { - lean_dec_ref(x_171); - x_201 = lean_box(0); + lean_dec_ref(x_182); + x_219 = lean_box(0); } -if (lean_is_scalar(x_201)) { - x_202 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_219)) { + x_220 = lean_alloc_ctor(1, 2, 0); } else { - x_202 = x_201; + x_220 = x_219; } -lean_ctor_set(x_202, 0, x_199); -lean_ctor_set(x_202, 1, x_200); -return x_202; +lean_ctor_set(x_220, 0, x_217); +lean_ctor_set(x_220, 1, x_218); +return x_220; } } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_203 = lean_ctor_get(x_168, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_168, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_168)) { - lean_ctor_release(x_168, 0); - lean_ctor_release(x_168, 1); - x_205 = x_168; +x_221 = lean_ctor_get(x_177, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_177, 1); +lean_inc(x_222); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_223 = x_177; } else { - lean_dec_ref(x_168); - x_205 = lean_box(0); + lean_dec_ref(x_177); + x_223 = lean_box(0); } -if (lean_is_scalar(x_205)) { - x_206 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_223)) { + x_224 = lean_alloc_ctor(1, 2, 0); } else { - x_206 = x_205; + x_224 = x_223; } -lean_ctor_set(x_206, 0, x_203); -lean_ctor_set(x_206, 1, x_204); -return x_206; +lean_ctor_set(x_224, 0, x_221); +lean_ctor_set(x_224, 1, x_222); +return x_224; } } else { -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_207 = lean_ctor_get(x_164, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_164, 1); -lean_inc(x_208); -if (lean_is_exclusive(x_164)) { - lean_ctor_release(x_164, 0); - lean_ctor_release(x_164, 1); - x_209 = x_164; +x_225 = lean_ctor_get(x_173, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_173, 1); +lean_inc(x_226); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_227 = x_173; } else { - lean_dec_ref(x_164); - x_209 = lean_box(0); + lean_dec_ref(x_173); + x_227 = lean_box(0); } -if (lean_is_scalar(x_209)) { - x_210 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_227)) { + x_228 = lean_alloc_ctor(1, 2, 0); } else { - x_210 = x_209; + x_228 = x_227; } -lean_ctor_set(x_210, 0, x_207); -lean_ctor_set(x_210, 1, x_208); -return x_210; +lean_ctor_set(x_228, 0, x_225); +lean_ctor_set(x_228, 1, x_226); +return x_228; } } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_211 = lean_ctor_get(x_158, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_158, 1); -lean_inc(x_212); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_213 = x_158; +x_229 = lean_ctor_get(x_167, 0); +lean_inc(x_229); +x_230 = lean_ctor_get(x_167, 1); +lean_inc(x_230); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_231 = x_167; } else { - lean_dec_ref(x_158); - x_213 = lean_box(0); + lean_dec_ref(x_167); + x_231 = lean_box(0); } -if (lean_is_scalar(x_213)) { - x_214 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_231)) { + x_232 = lean_alloc_ctor(1, 2, 0); } else { - x_214 = x_213; + x_232 = x_231; } -lean_ctor_set(x_214, 0, x_211); -lean_ctor_set(x_214, 1, x_212); -return x_214; +lean_ctor_set(x_232, 0, x_229); +lean_ctor_set(x_232, 1, x_230); +return x_232; } } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_215 = lean_ctor_get(x_153, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_153, 1); -lean_inc(x_216); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_217 = x_153; +x_233 = lean_ctor_get(x_162, 0); +lean_inc(x_233); +x_234 = lean_ctor_get(x_162, 1); +lean_inc(x_234); +if (lean_is_exclusive(x_162)) { + lean_ctor_release(x_162, 0); + lean_ctor_release(x_162, 1); + x_235 = x_162; } else { - lean_dec_ref(x_153); - x_217 = lean_box(0); + lean_dec_ref(x_162); + x_235 = lean_box(0); } -if (lean_is_scalar(x_217)) { - x_218 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_235)) { + x_236 = lean_alloc_ctor(1, 2, 0); } else { - x_218 = x_217; + x_236 = x_235; } -lean_ctor_set(x_218, 0, x_215); -lean_ctor_set(x_218, 1, x_216); -return x_218; +lean_ctor_set(x_236, 0, x_233); +lean_ctor_set(x_236, 1, x_234); +return x_236; } } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_219 = lean_ctor_get(x_148, 0); -lean_inc(x_219); -x_220 = lean_ctor_get(x_148, 1); -lean_inc(x_220); -if (lean_is_exclusive(x_148)) { - lean_ctor_release(x_148, 0); - lean_ctor_release(x_148, 1); - x_221 = x_148; +x_237 = lean_ctor_get(x_157, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_157, 1); +lean_inc(x_238); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_239 = x_157; } else { - lean_dec_ref(x_148); - x_221 = lean_box(0); + lean_dec_ref(x_157); + x_239 = lean_box(0); } -if (lean_is_scalar(x_221)) { - x_222 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_239)) { + x_240 = lean_alloc_ctor(1, 2, 0); } else { - x_222 = x_221; + x_240 = x_239; } -lean_ctor_set(x_222, 0, x_219); -lean_ctor_set(x_222, 1, x_220); -return x_222; +lean_ctor_set(x_240, 0, x_237); +lean_ctor_set(x_240, 1, x_238); +return x_240; } } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; -lean_dec(x_143); -lean_dec(x_137); +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; +lean_dec(x_152); +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_223 = lean_ctor_get(x_145, 0); -lean_inc(x_223); -x_224 = lean_ctor_get(x_145, 1); -lean_inc(x_224); -if (lean_is_exclusive(x_145)) { - lean_ctor_release(x_145, 0); - lean_ctor_release(x_145, 1); - x_225 = x_145; +x_241 = lean_ctor_get(x_154, 0); +lean_inc(x_241); +x_242 = lean_ctor_get(x_154, 1); +lean_inc(x_242); +if (lean_is_exclusive(x_154)) { + lean_ctor_release(x_154, 0); + lean_ctor_release(x_154, 1); + x_243 = x_154; } else { - lean_dec_ref(x_145); - x_225 = lean_box(0); + lean_dec_ref(x_154); + x_243 = lean_box(0); } -if (lean_is_scalar(x_225)) { - x_226 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_243)) { + x_244 = lean_alloc_ctor(1, 2, 0); } else { - x_226 = x_225; + x_244 = x_243; } -lean_ctor_set(x_226, 0, x_223); -lean_ctor_set(x_226, 1, x_224); -return x_226; +lean_ctor_set(x_244, 0, x_241); +lean_ctor_set(x_244, 1, x_242); +return x_244; } } else { -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; -lean_dec(x_137); +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +lean_dec(x_146); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_227 = lean_ctor_get(x_139, 0); -lean_inc(x_227); -x_228 = lean_ctor_get(x_139, 1); -lean_inc(x_228); -if (lean_is_exclusive(x_139)) { - lean_ctor_release(x_139, 0); - lean_ctor_release(x_139, 1); - x_229 = x_139; +x_245 = lean_ctor_get(x_148, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_148, 1); +lean_inc(x_246); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_247 = x_148; } else { - lean_dec_ref(x_139); - x_229 = lean_box(0); + lean_dec_ref(x_148); + x_247 = lean_box(0); } -if (lean_is_scalar(x_229)) { - x_230 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_247)) { + x_248 = lean_alloc_ctor(1, 2, 0); } else { - x_230 = x_229; + x_248 = x_247; } -lean_ctor_set(x_230, 0, x_227); -lean_ctor_set(x_230, 1, x_228); -return x_230; +lean_ctor_set(x_248, 0, x_245); +lean_ctor_set(x_248, 1, x_246); +return x_248; } } } @@ -1937,6 +2049,7 @@ lean_object* initialize_Lean_Meta_Tactic_Assert(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Simp_Main(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Util(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_DoNotSimp(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Grind_MarkNestedProofs(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object* w) { @@ -1958,6 +2071,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_DoNotSimp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Grind_MarkNestedProofs(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -1992,6 +2108,10 @@ l_Lean_Meta_Grind_simp___closed__12 = _init_l_Lean_Meta_Grind_simp___closed__12( lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__12); l_Lean_Meta_Grind_simp___closed__13 = _init_l_Lean_Meta_Grind_simp___closed__13(); lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__13); +l_Lean_Meta_Grind_simp___closed__14 = _init_l_Lean_Meta_Grind_simp___closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__14); +l_Lean_Meta_Grind_simp___closed__15 = _init_l_Lean_Meta_Grind_simp___closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_simp___closed__15); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/SimpUtil.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/SimpUtil.c new file mode 100644 index 000000000000..5a2641e603b2 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/SimpUtil.c @@ -0,0 +1,586 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.SimpUtil +// Imports: Lean.Meta.Tactic.Simp.Simproc Lean.Meta.Tactic.Grind.Simp Lean.Meta.Tactic.Grind.DoNotSimp +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__8; +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__6; +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_getSimpContext___closed__1; +lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimpContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__5; +lean_object* l_Lean_Meta_Grind_addDoNotSimp(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__4; +lean_object* l_Lean_Meta_Simp_SimprocExtension_getSimprocs(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__3; +lean_object* l_Lean_Meta_Simp_getSEvalSimprocs___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimpContext___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs(lean_object*, lean_object*); +extern lean_object* l_Lean_Meta_Grind_grindNormSimprocExt; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_getSimpContext___closed__2; +extern lean_object* l_Lean_Meta_Grind_grindNormExt; +LEAN_EXPORT lean_object* lean_grind_normalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_getSimprocs___rarg___closed__1; +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__1; +static lean_object* l_Lean_Meta_Grind_normalizeImp___closed__2; +static lean_object* _init_l_Lean_Meta_Grind_getSimprocs___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_grindNormSimprocExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = l_Lean_Meta_Grind_getSimprocs___rarg___closed__1; +x_5 = l_Lean_Meta_Simp_SimprocExtension_getSimprocs(x_4, x_1, x_2, x_3); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = l_Lean_Meta_Grind_addDoNotSimp(x_7, x_1, x_2, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Meta_Simp_getSEvalSimprocs___rarg(x_2, x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_box(0); +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 1, x_15); +lean_ctor_set(x_5, 0, x_14); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_5); +x_17 = lean_array_mk(x_16); +lean_ctor_set(x_12, 0, x_17); +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_ctor_get(x_12, 0); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_12); +x_20 = lean_box(0); +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 1, x_20); +lean_ctor_set(x_5, 0, x_18); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_5); +x_22 = lean_array_mk(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_19); +return x_23; +} +} +else +{ +uint8_t x_24; +lean_free_object(x_5); +x_24 = !lean_is_exclusive(x_9); +if (x_24 == 0) +{ +return x_9; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_9, 0); +x_26 = lean_ctor_get(x_9, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_9); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_5, 0); +x_29 = lean_ctor_get(x_5, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_5); +x_30 = l_Lean_Meta_Grind_addDoNotSimp(x_28, x_1, x_2, x_29); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Meta_Simp_getSEvalSimprocs___rarg(x_2, x_32); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + x_36 = x_33; +} else { + lean_dec_ref(x_33); + x_36 = lean_box(0); +} +x_37 = lean_box(0); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_37); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_31); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_array_mk(x_39); +if (lean_is_scalar(x_36)) { + x_41 = lean_alloc_ctor(0, 2, 0); +} else { + x_41 = x_36; +} +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_35); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_30, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_30, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_44 = x_30; +} else { + lean_dec_ref(x_30); + x_44 = lean_box(0); +} +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(1, 2, 0); +} else { + x_45 = x_44; +} +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +return x_45; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_getSimprocs___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_getSimprocs___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimprocs___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_getSimprocs(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_getSimpContext___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_grindNormExt; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_getSimpContext___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = l_Lean_Meta_Simp_defaultMaxSteps; +x_2 = lean_unsigned_to_nat(2u); +x_3 = 0; +x_4 = 1; +x_5 = 0; +x_6 = lean_alloc_ctor(0, 2, 19); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set_uint8(x_6, sizeof(void*)*2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 1, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 3, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 4, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 5, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 6, x_5); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 7, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 8, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 9, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 10, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 11, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 12, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 13, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 14, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 15, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 16, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 17, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*2 + 18, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimpContext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_6 = l_Lean_Meta_Grind_getSimpContext___closed__1; +x_7 = l_Lean_Meta_SimpExtension_getTheorems(x_6, x_3, x_4, x_5); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_4, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = lean_box(0); +lean_ctor_set_tag(x_10, 1); +lean_ctor_set(x_10, 1, x_14); +lean_ctor_set(x_10, 0, x_8); +x_15 = lean_array_mk(x_10); +x_16 = l_Lean_Meta_Grind_getSimpContext___closed__2; +x_17 = l_Lean_Meta_Simp_mkContext(x_16, x_15, x_12, x_1, x_2, x_3, x_4, x_13); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_10, 0); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_10); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_8); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_array_mk(x_21); +x_23 = l_Lean_Meta_Grind_getSimpContext___closed__2; +x_24 = l_Lean_Meta_Simp_mkContext(x_23, x_22, x_18, x_1, x_2, x_3, x_4, x_19); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getSimpContext___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_Meta_Grind_getSimpContext(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_normalizeImp___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_normalizeImp___closed__2; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_normalizeImp___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__6() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Meta_Grind_normalizeImp___closed__5; +x_3 = l_Lean_Meta_Grind_normalizeImp___closed__4; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_normalizeImp___closed__2; +x_2 = l_Lean_Meta_Grind_normalizeImp___closed__6; +x_3 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_normalizeImp___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_normalizeImp___closed__3; +x_2 = l_Lean_Meta_Grind_normalizeImp___closed__7; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* lean_grind_normalize(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = l_Lean_Meta_Grind_getSimpContext(x_2, x_3, x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = l_Lean_Meta_Grind_getSimprocs___rarg(x_4, x_5, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_box(0); +x_14 = l_Lean_Meta_Grind_normalizeImp___closed__8; +x_15 = l_Lean_Meta_simp(x_1, x_8, x_11, x_13, x_14, x_2, x_3, x_4, x_5, x_12); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_15, 0); +lean_dec(x_19); +x_20 = lean_ctor_get(x_17, 0); +lean_inc(x_20); +lean_dec(x_17); +lean_ctor_set(x_15, 0, x_20); +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +x_22 = lean_ctor_get(x_17, 0); +lean_inc(x_22); +lean_dec(x_17); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_15); +if (x_24 == 0) +{ +return x_15; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_15, 0); +x_26 = lean_ctor_get(x_15, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_15); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +else +{ +uint8_t x_28; +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_10); +if (x_28 == 0) +{ +return x_10; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_10, 0); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_10); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +lean_object* initialize_Lean_Meta_Tactic_Simp_Simproc(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Simp(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_DoNotSimp(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_SimpUtil(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Meta_Tactic_Simp_Simproc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Simp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_DoNotSimp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_getSimprocs___rarg___closed__1 = _init_l_Lean_Meta_Grind_getSimprocs___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_getSimprocs___rarg___closed__1); +l_Lean_Meta_Grind_getSimpContext___closed__1 = _init_l_Lean_Meta_Grind_getSimpContext___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_getSimpContext___closed__1); +l_Lean_Meta_Grind_getSimpContext___closed__2 = _init_l_Lean_Meta_Grind_getSimpContext___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_getSimpContext___closed__2); +l_Lean_Meta_Grind_normalizeImp___closed__1 = _init_l_Lean_Meta_Grind_normalizeImp___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__1); +l_Lean_Meta_Grind_normalizeImp___closed__2 = _init_l_Lean_Meta_Grind_normalizeImp___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__2); +l_Lean_Meta_Grind_normalizeImp___closed__3 = _init_l_Lean_Meta_Grind_normalizeImp___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__3); +l_Lean_Meta_Grind_normalizeImp___closed__4 = _init_l_Lean_Meta_Grind_normalizeImp___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__4); +l_Lean_Meta_Grind_normalizeImp___closed__5 = _init_l_Lean_Meta_Grind_normalizeImp___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__5); +l_Lean_Meta_Grind_normalizeImp___closed__6 = _init_l_Lean_Meta_Grind_normalizeImp___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__6); +l_Lean_Meta_Grind_normalizeImp___closed__7 = _init_l_Lean_Meta_Grind_normalizeImp___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__7); +l_Lean_Meta_Grind_normalizeImp___closed__8 = _init_l_Lean_Meta_Grind_normalizeImp___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_normalizeImp___closed__8); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Split.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Split.c new file mode 100644 index 000000000000..87eff9d04c71 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Split.c @@ -0,0 +1,5008 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Grind.Split +// Imports: Lean.Meta.Tactic.Grind.Types Lean.Meta.Tactic.Grind.Intro Lean.Meta.Tactic.Grind.Cases Lean.Meta.Tactic.Grind.CasesMatch +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3; +lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2; +lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_introNewHyp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_splitNext___lambda__4___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCaseSplitStatus; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1; +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1(lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18____boxed(lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofFormat(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_isInductivePredicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_casesMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1; +lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Meta_instMonadMetaM; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instInhabitedCaseSplitStatus; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6; +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_isResolvedCaseSplit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4; +lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_intros(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2; +lean_object* l_Lean_Expr_getAppFn(lean_object*); +lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5; +lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkEM(lean_object*); +lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion(lean_object*); +lean_object* l_Lean_Meta_Grind_isEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_splitNext___lambda__4___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_cases(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9; +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_splitNext___lambda__4___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_splitNext___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1___boxed(lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Grind_isInconsistent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5; +lean_object* l_ReaderT_instMonad___rarg(lean_object*); +static lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3; +lean_object* l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +default: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static uint8_t _init_l_Lean_Meta_Grind_instInhabitedCaseSplitStatus() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18_(uint8_t x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx(x_1); +x_4 = l_Lean_Meta_Grind_CaseSplitStatus_toCtorIdx(x_2); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = lean_unbox(x_2); +lean_dec(x_2); +x_5 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18_(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_beqCaseSplitStatus____x40_Lean_Meta_Tactic_Grind_Split___hyg_18____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instBEqCaseSplitStatus() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1; +return x_1; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMonadMetaM; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3; +x_2 = l_ReaderT_instMonad___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4; +x_2 = l_Lean_Meta_Grind_instInhabitedCaseSplitStatus; +x_3 = lean_box(x_2); +x_4 = l_instInhabitedOfMonad___rarg(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_Meta_Grind_isEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_15, 0); +lean_dec(x_19); +x_20 = 2; +x_21 = lean_box(x_20); +lean_ctor_set(x_15, 0, x_21); +return x_15; +} +else +{ +lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_dec(x_15); +x_23 = 2; +x_24 = lean_box(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_22); +return x_25; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_15); +if (x_26 == 0) +{ +lean_object* x_27; uint8_t x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_15, 0); +lean_dec(x_27); +x_28 = 0; +x_29 = lean_box(x_28); +lean_ctor_set(x_15, 0, x_29); +return x_15; +} +else +{ +lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_15, 1); +lean_inc(x_30); +lean_dec(x_15); +x_31 = 0; +x_32 = lean_box(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_30); +return x_33; +} +} +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) +{ +return x_15; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_ctor_get(x_15, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_15); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +uint8_t x_38; +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_11); +if (x_38 == 0) +{ +lean_object* x_39; uint8_t x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_11, 0); +lean_dec(x_39); +x_40 = 0; +x_41 = lean_box(x_40); +lean_ctor_set(x_11, 0, x_41); +return x_11; +} +else +{ +lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_11, 1); +lean_inc(x_42); +lean_dec(x_11); +x_43 = 0; +x_44 = lean_box(x_43); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_42); +return x_45; +} +} +} +else +{ +uint8_t x_46; +lean_dec(x_1); +x_46 = !lean_is_exclusive(x_11); +if (x_46 == 0) +{ +return x_11; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_11, 0); +x_48 = lean_ctor_get(x_11, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_11); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_isEqTrue(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l_Lean_Meta_Grind_isEqFalse(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; +lean_dec(x_3); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_17, 0, x_23); +return x_17; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_dec(x_17); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_17, 1); +lean_inc(x_28); +lean_dec(x_17); +x_29 = l_Lean_Meta_Grind_isEqFalse(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; uint8_t x_31; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_unbox(x_30); +lean_dec(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = l_Lean_Meta_Grind_isEqFalse(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_32); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; uint8_t x_35; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_unbox(x_34); +lean_dec(x_34); +if (x_35 == 0) +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_33); +if (x_36 == 0) +{ +lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 0); +lean_dec(x_37); +x_38 = 2; +x_39 = lean_box(x_38); +lean_ctor_set(x_33, 0, x_39); +return x_33; +} +else +{ +lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_dec(x_33); +x_41 = 2; +x_42 = lean_box(x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; +} +} +else +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_33); +if (x_44 == 0) +{ +lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_33, 0); +lean_dec(x_45); +x_46 = 0; +x_47 = lean_box(x_46); +lean_ctor_set(x_33, 0, x_47); +return x_33; +} +else +{ +lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; +x_48 = lean_ctor_get(x_33, 1); +lean_inc(x_48); +lean_dec(x_33); +x_49 = 0; +x_50 = lean_box(x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; +} +} +} +else +{ +uint8_t x_52; +x_52 = !lean_is_exclusive(x_33); +if (x_52 == 0) +{ +return x_33; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_33, 0); +x_54 = lean_ctor_get(x_33, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_33); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +else +{ +uint8_t x_56; +lean_dec(x_3); +x_56 = !lean_is_exclusive(x_29); +if (x_56 == 0) +{ +lean_object* x_57; uint8_t x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_29, 0); +lean_dec(x_57); +x_58 = 0; +x_59 = lean_box(x_58); +lean_ctor_set(x_29, 0, x_59); +return x_29; +} +else +{ +lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_29, 1); +lean_inc(x_60); +lean_dec(x_29); +x_61 = 0; +x_62 = lean_box(x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_60); +return x_63; +} +} +} +else +{ +uint8_t x_64; +lean_dec(x_3); +x_64 = !lean_is_exclusive(x_29); +if (x_64 == 0) +{ +return x_29; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_29, 0); +x_66 = lean_ctor_get(x_29, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_29); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +} +else +{ +uint8_t x_68; +lean_dec(x_3); +lean_dec(x_2); +x_68 = !lean_is_exclusive(x_17); +if (x_68 == 0) +{ +return x_17; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_17, 0); +x_70 = lean_ctor_get(x_17, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_17); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +else +{ +uint8_t x_72; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_72 = !lean_is_exclusive(x_13); +if (x_72 == 0) +{ +lean_object* x_73; uint8_t x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_13, 0); +lean_dec(x_73); +x_74 = 0; +x_75 = lean_box(x_74); +lean_ctor_set(x_13, 0, x_75); +return x_13; +} +else +{ +lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_13, 1); +lean_inc(x_76); +lean_dec(x_13); +x_77 = 0; +x_78 = lean_box(x_77); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_76); +return x_79; +} +} +} +else +{ +uint8_t x_80; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_80 = !lean_is_exclusive(x_13); +if (x_80 == 0) +{ +return x_13; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_13, 0); +x_82 = lean_ctor_get(x_13, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_13); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_isEqTrue(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l_Lean_Meta_Grind_isEqFalse(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_17, 0, x_23); +return x_17; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_dec(x_17); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_17); +if (x_28 == 0) +{ +lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_17, 0); +lean_dec(x_29); +x_30 = 0; +x_31 = lean_box(x_30); +lean_ctor_set(x_17, 0, x_31); +return x_17; +} +else +{ +lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_17, 1); +lean_inc(x_32); +lean_dec(x_17); +x_33 = 0; +x_34 = lean_box(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_32); +return x_35; +} +} +} +else +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_17); +if (x_36 == 0) +{ +return x_17; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_17, 0); +x_38 = lean_ctor_get(x_17, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_17); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +lean_object* x_40; lean_object* x_41; +lean_dec(x_1); +x_40 = lean_ctor_get(x_13, 1); +lean_inc(x_40); +lean_dec(x_13); +x_41 = l_Lean_Meta_Grind_isEqTrue(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; uint8_t x_43; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_unbox(x_42); +lean_dec(x_42); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_41, 1); +lean_inc(x_44); +lean_dec(x_41); +x_45 = l_Lean_Meta_Grind_isEqTrue(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; uint8_t x_47; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) +{ +uint8_t x_48; +x_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) +{ +lean_object* x_49; uint8_t x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 0); +lean_dec(x_49); +x_50 = 2; +x_51 = lean_box(x_50); +lean_ctor_set(x_45, 0, x_51); +return x_45; +} +else +{ +lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_45, 1); +lean_inc(x_52); +lean_dec(x_45); +x_53 = 2; +x_54 = lean_box(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_52); +return x_55; +} +} +else +{ +uint8_t x_56; +x_56 = !lean_is_exclusive(x_45); +if (x_56 == 0) +{ +lean_object* x_57; uint8_t x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_45, 0); +lean_dec(x_57); +x_58 = 0; +x_59 = lean_box(x_58); +lean_ctor_set(x_45, 0, x_59); +return x_45; +} +else +{ +lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_45, 1); +lean_inc(x_60); +lean_dec(x_45); +x_61 = 0; +x_62 = lean_box(x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_60); +return x_63; +} +} +} +else +{ +uint8_t x_64; +x_64 = !lean_is_exclusive(x_45); +if (x_64 == 0) +{ +return x_45; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_45, 0); +x_66 = lean_ctor_get(x_45, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_45); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +uint8_t x_68; +lean_dec(x_3); +x_68 = !lean_is_exclusive(x_41); +if (x_68 == 0) +{ +lean_object* x_69; uint8_t x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_41, 0); +lean_dec(x_69); +x_70 = 0; +x_71 = lean_box(x_70); +lean_ctor_set(x_41, 0, x_71); +return x_41; +} +else +{ +lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; +x_72 = lean_ctor_get(x_41, 1); +lean_inc(x_72); +lean_dec(x_41); +x_73 = 0; +x_74 = lean_box(x_73); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_72); +return x_75; +} +} +} +else +{ +uint8_t x_76; +lean_dec(x_3); +x_76 = !lean_is_exclusive(x_41); +if (x_76 == 0) +{ +return x_41; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_41, 0); +x_78 = lean_ctor_get(x_41, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_41); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} +} +} +else +{ +uint8_t x_80; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_80 = !lean_is_exclusive(x_13); +if (x_80 == 0) +{ +return x_13; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_13, 0); +x_82 = lean_ctor_get(x_13, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_13); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_11 = 1; +x_12 = lean_box(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +return x_13; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Meta.Tactic.Grind.Split", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Meta.Tactic.Grind.Split.0.Lean.Meta.Grind.checkCaseSplitStatus", 76, 76); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2; +x_3 = lean_unsigned_to_nat(58u); +x_4 = lean_unsigned_to_nat(43u); +x_5 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4___boxed), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_12) == 4) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_14 = l_Lean_Meta_isInductivePredicate(x_13, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_unbox(x_15); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_1); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5; +x_19 = lean_box(0); +x_20 = lean_apply_10(x_18, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_14, 1); +lean_inc(x_21); +lean_dec(x_14); +x_22 = l_Lean_Meta_Grind_isEqTrue(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5; +x_27 = lean_box(0); +x_28 = lean_apply_10(x_26, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +return x_28; +} +else +{ +uint8_t x_29; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_29 = !lean_is_exclusive(x_22); +if (x_29 == 0) +{ +lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_22, 0); +lean_dec(x_30); +x_31 = 2; +x_32 = lean_box(x_31); +lean_ctor_set(x_22, 0, x_32); +return x_22; +} +else +{ +lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_22, 1); +lean_inc(x_33); +lean_dec(x_22); +x_34 = 2; +x_35 = lean_box(x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_33); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_37 = !lean_is_exclusive(x_22); +if (x_37 == 0) +{ +return x_22; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_22, 0); +x_39 = lean_ctor_get(x_22, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_22); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +} +else +{ +uint8_t x_41; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_41 = !lean_is_exclusive(x_14); +if (x_41 == 0) +{ +return x_14; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_14, 0); +x_43 = lean_ctor_get(x_14, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_14); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +} +else +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_12); +lean_dec(x_1); +x_45 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4; +x_46 = l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1(x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_46; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = 2; +x_21 = lean_box(x_20); +lean_ctor_set(x_12, 0, x_21); +return x_12; +} +else +{ +lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_dec(x_12); +x_23 = 2; +x_24 = lean_box(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_22); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_box(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +return x_13; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("debug", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7___boxed), 10, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split resolved: ", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_isResolvedCaseSplit(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} +else +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_12, 1); +x_20 = lean_ctor_get(x_12, 0); +lean_dec(x_20); +x_21 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4; +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_ctor_get(x_22, 1); +x_26 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5; +x_27 = lean_unbox(x_24); +lean_dec(x_24); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +lean_free_object(x_22); +lean_free_object(x_12); +lean_dec(x_1); +x_28 = lean_box(0); +x_29 = lean_apply_10(x_26, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_30 = l_Lean_MessageData_ofExpr(x_1); +x_31 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7; +lean_ctor_set_tag(x_22, 7); +lean_ctor_set(x_22, 1, x_30); +lean_ctor_set(x_22, 0, x_31); +x_32 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_32); +lean_ctor_set(x_12, 0, x_22); +x_33 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_21, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_apply_10(x_26, x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +return x_36; +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_37 = lean_ctor_get(x_22, 0); +x_38 = lean_ctor_get(x_22, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_22); +x_39 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5; +x_40 = lean_unbox(x_37); +lean_dec(x_37); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +lean_free_object(x_12); +lean_dec(x_1); +x_41 = lean_box(0); +x_42 = lean_apply_10(x_39, x_41, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_38); +return x_42; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_43 = l_Lean_MessageData_ofExpr(x_1); +x_44 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +x_46 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +lean_ctor_set_tag(x_12, 7); +lean_ctor_set(x_12, 1, x_46); +lean_ctor_set(x_12, 0, x_45); +x_47 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_21, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_38); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = lean_apply_10(x_39, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_49); +return x_50; +} +} +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_51 = lean_ctor_get(x_12, 1); +lean_inc(x_51); +lean_dec(x_12); +x_52 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4; +x_53 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_51); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_56 = x_53; +} else { + lean_dec_ref(x_53); + x_56 = lean_box(0); +} +x_57 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5; +x_58 = lean_unbox(x_54); +lean_dec(x_54); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_56); +lean_dec(x_1); +x_59 = lean_box(0); +x_60 = lean_apply_10(x_57, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_55); +return x_60; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_61 = l_Lean_MessageData_ofExpr(x_1); +x_62 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7; +if (lean_is_scalar(x_56)) { + x_63 = lean_alloc_ctor(7, 2, 0); +} else { + x_63 = x_56; + lean_ctor_set_tag(x_63, 7); +} +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +x_64 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_52, x_65, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_55); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_apply_10(x_57, x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_68); +return x_69; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1___boxed), 10, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("And", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Or", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dite", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ite", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_inc(x_1); +x_11 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1; +x_15 = l_Lean_Expr_cleanupAnnotations(x_12); +x_16 = l_Lean_Expr_isApp(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_15); +x_17 = lean_box(0); +x_18 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Expr_appArg(x_15, lean_box(0)); +x_20 = l_Lean_Expr_appFnCleanup(x_15, lean_box(0)); +x_21 = l_Lean_Expr_isApp(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_19); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); +x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); +x_26 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3; +x_27 = l_Lean_Expr_isConstOf(x_25, x_26); +if (x_27 == 0) +{ +lean_object* x_28; uint8_t x_29; +x_28 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5; +x_29 = l_Lean_Expr_isConstOf(x_25, x_28); +if (x_29 == 0) +{ +uint8_t x_30; +lean_dec(x_24); +lean_dec(x_19); +x_30 = l_Lean_Expr_isApp(x_25); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_25); +x_31 = lean_box(0); +x_32 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_32; +} +else +{ +lean_object* x_33; uint8_t x_34; +x_33 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); +x_34 = l_Lean_Expr_isApp(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_33); +x_35 = lean_box(0); +x_36 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = l_Lean_Expr_appArg(x_33, lean_box(0)); +x_38 = l_Lean_Expr_appFnCleanup(x_33, lean_box(0)); +x_39 = l_Lean_Expr_isApp(x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +lean_dec(x_38); +lean_dec(x_37); +x_40 = lean_box(0); +x_41 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = l_Lean_Expr_appFnCleanup(x_38, lean_box(0)); +x_43 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7; +x_44 = l_Lean_Expr_isConstOf(x_42, x_43); +if (x_44 == 0) +{ +lean_object* x_45; uint8_t x_46; +x_45 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9; +x_46 = l_Lean_Expr_isConstOf(x_42, x_45); +lean_dec(x_42); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; +lean_dec(x_37); +x_47 = lean_box(0); +x_48 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_48; +} +else +{ +lean_object* x_49; +lean_dec(x_1); +x_49 = lean_apply_10(x_14, x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_49; +} +} +else +{ +lean_object* x_50; +lean_dec(x_42); +lean_dec(x_1); +x_50 = lean_apply_10(x_14, x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_50; +} +} +} +} +} +else +{ +lean_object* x_51; +lean_dec(x_25); +x_51 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3(x_1, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_51; +} +} +else +{ +lean_object* x_52; +lean_dec(x_25); +x_52 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2(x_1, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_52; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_13 = lean_st_ref_take(x_4, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_14, 16); +lean_dec(x_17); +x_18 = l_List_reverse___rarg(x_3); +lean_ctor_set(x_14, 16, x_18); +x_19 = lean_st_ref_set(x_4, x_14, x_15); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_20; +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_2); +return x_19; +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_2); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_st_ref_take(x_4, x_24); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_29 = lean_ctor_get(x_26, 17); +x_30 = lean_ctor_get(x_26, 12); +lean_dec(x_30); +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_add(x_29, x_31); +lean_dec(x_29); +x_33 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_26, 17, x_32); +lean_ctor_set(x_26, 12, x_33); +x_34 = lean_st_ref_set(x_4, x_26, x_27); +lean_dec(x_4); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_34, 0); +lean_dec(x_36); +lean_ctor_set(x_34, 0, x_2); +return x_34; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_2); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_39 = lean_ctor_get(x_26, 0); +x_40 = lean_ctor_get(x_26, 1); +x_41 = lean_ctor_get(x_26, 2); +x_42 = lean_ctor_get(x_26, 3); +x_43 = lean_ctor_get(x_26, 4); +x_44 = lean_ctor_get(x_26, 5); +x_45 = lean_ctor_get_uint8(x_26, sizeof(void*)*19); +x_46 = lean_ctor_get(x_26, 6); +x_47 = lean_ctor_get(x_26, 7); +x_48 = lean_ctor_get(x_26, 8); +x_49 = lean_ctor_get(x_26, 9); +x_50 = lean_ctor_get(x_26, 10); +x_51 = lean_ctor_get(x_26, 11); +x_52 = lean_ctor_get(x_26, 13); +x_53 = lean_ctor_get(x_26, 14); +x_54 = lean_ctor_get(x_26, 15); +x_55 = lean_ctor_get(x_26, 16); +x_56 = lean_ctor_get(x_26, 17); +x_57 = lean_ctor_get(x_26, 18); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_26); +x_58 = lean_unsigned_to_nat(1u); +x_59 = lean_nat_add(x_56, x_58); +lean_dec(x_56); +x_60 = lean_unsigned_to_nat(0u); +x_61 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_61, 0, x_39); +lean_ctor_set(x_61, 1, x_40); +lean_ctor_set(x_61, 2, x_41); +lean_ctor_set(x_61, 3, x_42); +lean_ctor_set(x_61, 4, x_43); +lean_ctor_set(x_61, 5, x_44); +lean_ctor_set(x_61, 6, x_46); +lean_ctor_set(x_61, 7, x_47); +lean_ctor_set(x_61, 8, x_48); +lean_ctor_set(x_61, 9, x_49); +lean_ctor_set(x_61, 10, x_50); +lean_ctor_set(x_61, 11, x_51); +lean_ctor_set(x_61, 12, x_60); +lean_ctor_set(x_61, 13, x_52); +lean_ctor_set(x_61, 14, x_53); +lean_ctor_set(x_61, 15, x_54); +lean_ctor_set(x_61, 16, x_55); +lean_ctor_set(x_61, 17, x_59); +lean_ctor_set(x_61, 18, x_57); +lean_ctor_set_uint8(x_61, sizeof(void*)*19, x_45); +x_62 = lean_st_ref_set(x_4, x_61, x_27); +lean_dec(x_4); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_64 = x_62; +} else { + lean_dec_ref(x_62); + x_64 = lean_box(0); +} +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_64; +} +lean_ctor_set(x_65, 0, x_2); +lean_ctor_set(x_65, 1, x_63); +return x_65; +} +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_66 = lean_ctor_get(x_14, 0); +x_67 = lean_ctor_get(x_14, 1); +x_68 = lean_ctor_get(x_14, 2); +x_69 = lean_ctor_get(x_14, 3); +x_70 = lean_ctor_get(x_14, 4); +x_71 = lean_ctor_get(x_14, 5); +x_72 = lean_ctor_get_uint8(x_14, sizeof(void*)*19); +x_73 = lean_ctor_get(x_14, 6); +x_74 = lean_ctor_get(x_14, 7); +x_75 = lean_ctor_get(x_14, 8); +x_76 = lean_ctor_get(x_14, 9); +x_77 = lean_ctor_get(x_14, 10); +x_78 = lean_ctor_get(x_14, 11); +x_79 = lean_ctor_get(x_14, 12); +x_80 = lean_ctor_get(x_14, 13); +x_81 = lean_ctor_get(x_14, 14); +x_82 = lean_ctor_get(x_14, 15); +x_83 = lean_ctor_get(x_14, 17); +x_84 = lean_ctor_get(x_14, 18); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_inc(x_79); +lean_inc(x_78); +lean_inc(x_77); +lean_inc(x_76); +lean_inc(x_75); +lean_inc(x_74); +lean_inc(x_73); +lean_inc(x_71); +lean_inc(x_70); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_14); +x_85 = l_List_reverse___rarg(x_3); +x_86 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_86, 0, x_66); +lean_ctor_set(x_86, 1, x_67); +lean_ctor_set(x_86, 2, x_68); +lean_ctor_set(x_86, 3, x_69); +lean_ctor_set(x_86, 4, x_70); +lean_ctor_set(x_86, 5, x_71); +lean_ctor_set(x_86, 6, x_73); +lean_ctor_set(x_86, 7, x_74); +lean_ctor_set(x_86, 8, x_75); +lean_ctor_set(x_86, 9, x_76); +lean_ctor_set(x_86, 10, x_77); +lean_ctor_set(x_86, 11, x_78); +lean_ctor_set(x_86, 12, x_79); +lean_ctor_set(x_86, 13, x_80); +lean_ctor_set(x_86, 14, x_81); +lean_ctor_set(x_86, 15, x_82); +lean_ctor_set(x_86, 16, x_85); +lean_ctor_set(x_86, 17, x_83); +lean_ctor_set(x_86, 18, x_84); +lean_ctor_set_uint8(x_86, sizeof(void*)*19, x_72); +x_87 = lean_st_ref_set(x_4, x_86, x_15); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_dec(x_4); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_89 = x_87; +} else { + lean_dec_ref(x_87); + x_89 = lean_box(0); +} +if (lean_is_scalar(x_89)) { + x_90 = lean_alloc_ctor(0, 2, 0); +} else { + x_90 = x_89; +} +lean_ctor_set(x_90, 0, x_2); +lean_ctor_set(x_90, 1, x_88); +return x_90; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_91 = lean_ctor_get(x_87, 1); +lean_inc(x_91); +lean_dec(x_87); +x_92 = lean_st_ref_take(x_4, x_91); +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_93, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_93, 1); +lean_inc(x_96); +x_97 = lean_ctor_get(x_93, 2); +lean_inc(x_97); +x_98 = lean_ctor_get(x_93, 3); +lean_inc(x_98); +x_99 = lean_ctor_get(x_93, 4); +lean_inc(x_99); +x_100 = lean_ctor_get(x_93, 5); +lean_inc(x_100); +x_101 = lean_ctor_get_uint8(x_93, sizeof(void*)*19); +x_102 = lean_ctor_get(x_93, 6); +lean_inc(x_102); +x_103 = lean_ctor_get(x_93, 7); +lean_inc(x_103); +x_104 = lean_ctor_get(x_93, 8); +lean_inc(x_104); +x_105 = lean_ctor_get(x_93, 9); +lean_inc(x_105); +x_106 = lean_ctor_get(x_93, 10); +lean_inc(x_106); +x_107 = lean_ctor_get(x_93, 11); +lean_inc(x_107); +x_108 = lean_ctor_get(x_93, 13); +lean_inc(x_108); +x_109 = lean_ctor_get(x_93, 14); +lean_inc(x_109); +x_110 = lean_ctor_get(x_93, 15); +lean_inc(x_110); +x_111 = lean_ctor_get(x_93, 16); +lean_inc(x_111); +x_112 = lean_ctor_get(x_93, 17); +lean_inc(x_112); +x_113 = lean_ctor_get(x_93, 18); +lean_inc(x_113); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + lean_ctor_release(x_93, 2); + lean_ctor_release(x_93, 3); + lean_ctor_release(x_93, 4); + lean_ctor_release(x_93, 5); + lean_ctor_release(x_93, 6); + lean_ctor_release(x_93, 7); + lean_ctor_release(x_93, 8); + lean_ctor_release(x_93, 9); + lean_ctor_release(x_93, 10); + lean_ctor_release(x_93, 11); + lean_ctor_release(x_93, 12); + lean_ctor_release(x_93, 13); + lean_ctor_release(x_93, 14); + lean_ctor_release(x_93, 15); + lean_ctor_release(x_93, 16); + lean_ctor_release(x_93, 17); + lean_ctor_release(x_93, 18); + x_114 = x_93; +} else { + lean_dec_ref(x_93); + x_114 = lean_box(0); +} +x_115 = lean_unsigned_to_nat(1u); +x_116 = lean_nat_add(x_112, x_115); +lean_dec(x_112); +x_117 = lean_unsigned_to_nat(0u); +if (lean_is_scalar(x_114)) { + x_118 = lean_alloc_ctor(0, 19, 1); +} else { + x_118 = x_114; +} +lean_ctor_set(x_118, 0, x_95); +lean_ctor_set(x_118, 1, x_96); +lean_ctor_set(x_118, 2, x_97); +lean_ctor_set(x_118, 3, x_98); +lean_ctor_set(x_118, 4, x_99); +lean_ctor_set(x_118, 5, x_100); +lean_ctor_set(x_118, 6, x_102); +lean_ctor_set(x_118, 7, x_103); +lean_ctor_set(x_118, 8, x_104); +lean_ctor_set(x_118, 9, x_105); +lean_ctor_set(x_118, 10, x_106); +lean_ctor_set(x_118, 11, x_107); +lean_ctor_set(x_118, 12, x_117); +lean_ctor_set(x_118, 13, x_108); +lean_ctor_set(x_118, 14, x_109); +lean_ctor_set(x_118, 15, x_110); +lean_ctor_set(x_118, 16, x_111); +lean_ctor_set(x_118, 17, x_116); +lean_ctor_set(x_118, 18, x_113); +lean_ctor_set_uint8(x_118, sizeof(void*)*19, x_101); +x_119 = lean_st_ref_set(x_4, x_118, x_94); +lean_dec(x_4); +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_121 = x_119; +} else { + lean_dec_ref(x_119); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(0, 2, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_2); +lean_ctor_set(x_122, 1, x_120); +return x_122; +} +} +} +else +{ +uint8_t x_123; +x_123 = !lean_is_exclusive(x_1); +if (x_123 == 0) +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_1, 0); +x_125 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_124); +x_126 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus(x_124, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_126) == 0) +{ +lean_object* x_127; uint8_t x_128; +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_unbox(x_127); +lean_dec(x_127); +switch (x_128) { +case 0: +{ +lean_object* x_129; +lean_free_object(x_1); +lean_dec(x_124); +x_129 = lean_ctor_get(x_126, 1); +lean_inc(x_129); +lean_dec(x_126); +x_1 = x_125; +x_12 = x_129; +goto _start; +} +case 1: +{ +lean_object* x_131; +x_131 = lean_ctor_get(x_126, 1); +lean_inc(x_131); +lean_dec(x_126); +lean_ctor_set(x_1, 1, x_3); +{ +lean_object* _tmp_0 = x_125; +lean_object* _tmp_2 = x_1; +lean_object* _tmp_11 = x_131; +x_1 = _tmp_0; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +default: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_133; lean_object* x_134; +lean_free_object(x_1); +x_133 = lean_ctor_get(x_126, 1); +lean_inc(x_133); +lean_dec(x_126); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_124); +x_1 = x_125; +x_2 = x_134; +x_12 = x_133; +goto _start; +} +else +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_126, 1); +lean_inc(x_136); +lean_dec(x_126); +x_137 = lean_ctor_get(x_2, 0); +lean_inc(x_137); +lean_inc(x_124); +x_138 = l_Lean_Meta_Grind_getGeneration(x_124, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_136); +if (lean_obj_tag(x_138) == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_138, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_138, 1); +lean_inc(x_140); +lean_dec(x_138); +lean_inc(x_137); +x_141 = l_Lean_Meta_Grind_getGeneration(x_137, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_140); +if (lean_obj_tag(x_141) == 0) +{ +lean_object* x_142; lean_object* x_143; uint8_t x_144; +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +lean_dec(x_141); +x_144 = lean_nat_dec_lt(x_139, x_142); +lean_dec(x_142); +lean_dec(x_139); +if (x_144 == 0) +{ +lean_dec(x_137); +lean_ctor_set(x_1, 1, x_3); +{ +lean_object* _tmp_0 = x_125; +lean_object* _tmp_2 = x_1; +lean_object* _tmp_11 = x_143; +x_1 = _tmp_0; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +else +{ +uint8_t x_146; +x_146 = !lean_is_exclusive(x_2); +if (x_146 == 0) +{ +lean_object* x_147; +x_147 = lean_ctor_get(x_2, 0); +lean_dec(x_147); +lean_ctor_set(x_2, 0, x_124); +lean_ctor_set(x_1, 1, x_3); +lean_ctor_set(x_1, 0, x_137); +{ +lean_object* _tmp_0 = x_125; +lean_object* _tmp_2 = x_1; +lean_object* _tmp_11 = x_143; +x_1 = _tmp_0; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +else +{ +lean_object* x_149; +lean_dec(x_2); +x_149 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_149, 0, x_124); +lean_ctor_set(x_1, 1, x_3); +lean_ctor_set(x_1, 0, x_137); +{ +lean_object* _tmp_0 = x_125; +lean_object* _tmp_1 = x_149; +lean_object* _tmp_2 = x_1; +lean_object* _tmp_11 = x_143; +x_1 = _tmp_0; +x_2 = _tmp_1; +x_3 = _tmp_2; +x_12 = _tmp_11; +} +goto _start; +} +} +} +else +{ +uint8_t x_151; +lean_dec(x_139); +lean_dec(x_137); +lean_free_object(x_1); +lean_dec(x_125); +lean_dec(x_124); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_151 = !lean_is_exclusive(x_141); +if (x_151 == 0) +{ +return x_141; +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_141, 0); +x_153 = lean_ctor_get(x_141, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_141); +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +return x_154; +} +} +} +else +{ +uint8_t x_155; +lean_dec(x_137); +lean_free_object(x_1); +lean_dec(x_125); +lean_dec(x_124); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_155 = !lean_is_exclusive(x_138); +if (x_155 == 0) +{ +return x_138; +} +else +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_138, 0); +x_157 = lean_ctor_get(x_138, 1); +lean_inc(x_157); +lean_inc(x_156); +lean_dec(x_138); +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +return x_158; +} +} +} +} +} +} +else +{ +uint8_t x_159; +lean_free_object(x_1); +lean_dec(x_125); +lean_dec(x_124); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_159 = !lean_is_exclusive(x_126); +if (x_159 == 0) +{ +return x_126; +} +else +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_126, 0); +x_161 = lean_ctor_get(x_126, 1); +lean_inc(x_161); +lean_inc(x_160); +lean_dec(x_126); +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +return x_162; +} +} +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get(x_1, 0); +x_164 = lean_ctor_get(x_1, 1); +lean_inc(x_164); +lean_inc(x_163); +lean_dec(x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_163); +x_165 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus(x_163, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; uint8_t x_167; +x_166 = lean_ctor_get(x_165, 0); +lean_inc(x_166); +x_167 = lean_unbox(x_166); +lean_dec(x_166); +switch (x_167) { +case 0: +{ +lean_object* x_168; +lean_dec(x_163); +x_168 = lean_ctor_get(x_165, 1); +lean_inc(x_168); +lean_dec(x_165); +x_1 = x_164; +x_12 = x_168; +goto _start; +} +case 1: +{ +lean_object* x_170; lean_object* x_171; +x_170 = lean_ctor_get(x_165, 1); +lean_inc(x_170); +lean_dec(x_165); +x_171 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_171, 0, x_163); +lean_ctor_set(x_171, 1, x_3); +x_1 = x_164; +x_3 = x_171; +x_12 = x_170; +goto _start; +} +default: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_173; lean_object* x_174; +x_173 = lean_ctor_get(x_165, 1); +lean_inc(x_173); +lean_dec(x_165); +x_174 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_174, 0, x_163); +x_1 = x_164; +x_2 = x_174; +x_12 = x_173; +goto _start; +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_176 = lean_ctor_get(x_165, 1); +lean_inc(x_176); +lean_dec(x_165); +x_177 = lean_ctor_get(x_2, 0); +lean_inc(x_177); +lean_inc(x_163); +x_178 = l_Lean_Meta_Grind_getGeneration(x_163, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_176); +if (lean_obj_tag(x_178) == 0) +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_178, 1); +lean_inc(x_180); +lean_dec(x_178); +lean_inc(x_177); +x_181 = l_Lean_Meta_Grind_getGeneration(x_177, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_180); +if (lean_obj_tag(x_181) == 0) +{ +lean_object* x_182; lean_object* x_183; uint8_t x_184; +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); +x_184 = lean_nat_dec_lt(x_179, x_182); +lean_dec(x_182); +lean_dec(x_179); +if (x_184 == 0) +{ +lean_object* x_185; +lean_dec(x_177); +x_185 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_185, 0, x_163); +lean_ctor_set(x_185, 1, x_3); +x_1 = x_164; +x_3 = x_185; +x_12 = x_183; +goto _start; +} +else +{ +lean_object* x_187; lean_object* x_188; lean_object* x_189; +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + x_187 = x_2; +} else { + lean_dec_ref(x_2); + x_187 = lean_box(0); +} +if (lean_is_scalar(x_187)) { + x_188 = lean_alloc_ctor(1, 1, 0); +} else { + x_188 = x_187; +} +lean_ctor_set(x_188, 0, x_163); +x_189 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_189, 0, x_177); +lean_ctor_set(x_189, 1, x_3); +x_1 = x_164; +x_2 = x_188; +x_3 = x_189; +x_12 = x_183; +goto _start; +} +} +else +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +lean_dec(x_179); +lean_dec(x_177); +lean_dec(x_164); +lean_dec(x_163); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_191 = lean_ctor_get(x_181, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_181, 1); +lean_inc(x_192); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_193 = x_181; +} else { + lean_dec_ref(x_181); + x_193 = lean_box(0); +} +if (lean_is_scalar(x_193)) { + x_194 = lean_alloc_ctor(1, 2, 0); +} else { + x_194 = x_193; +} +lean_ctor_set(x_194, 0, x_191); +lean_ctor_set(x_194, 1, x_192); +return x_194; +} +} +else +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_177); +lean_dec(x_164); +lean_dec(x_163); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_195 = lean_ctor_get(x_178, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_178, 1); +lean_inc(x_196); +if (lean_is_exclusive(x_178)) { + lean_ctor_release(x_178, 0); + lean_ctor_release(x_178, 1); + x_197 = x_178; +} else { + lean_dec_ref(x_178); + x_197 = lean_box(0); +} +if (lean_is_scalar(x_197)) { + x_198 = lean_alloc_ctor(1, 2, 0); +} else { + x_198 = x_197; +} +lean_ctor_set(x_198, 0, x_195); +lean_ctor_set(x_198, 1, x_196); +return x_198; +} +} +} +} +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_dec(x_164); +lean_dec(x_163); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_199 = lean_ctor_get(x_165, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_165, 1); +lean_inc(x_200); +if (lean_is_exclusive(x_165)) { + lean_ctor_release(x_165, 0); + lean_ctor_release(x_165, 1); + x_201 = x_165; +} else { + lean_dec_ref(x_165); + x_201 = lean_box(0); +} +if (lean_is_scalar(x_201)) { + x_202 = lean_alloc_ctor(1, 2, 0); +} else { + x_202 = x_201; +} +lean_ctor_set(x_202, 0, x_199); +lean_ctor_set(x_202, 1, x_200); +return x_202; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 16); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_box(0); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f_go(x_14, x_15, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_17; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1___boxed), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = l_Lean_Meta_Grind_checkMaxCaseSplit(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1; +x_16 = lean_box(0); +x_17 = lean_apply_10(x_15, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_18 = !lean_is_exclusive(x_11); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_11, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_11, 0, x_20); +return x_11; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_dec(x_11); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___boxed), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_Meta_Grind_isInconsistent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1; +x_15 = lean_box(0); +x_16 = lean_apply_10(x_14, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +return x_16; +} +else +{ +uint8_t x_17; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_10); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_10, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_10, 0, x_19); +return x_10; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_10, 1); +lean_inc(x_20); +lean_dec(x_10); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_mkEM(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Grind", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("or_of_and_eq_false", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_mkEqFalseProof(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5; +x_17 = l_Lean_mkApp3(x_16, x_2, x_3, x_15); +lean_ctor_set(x_13, 0, x_17); +return x_13; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_13, 0); +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_13); +x_20 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5; +x_21 = l_Lean_mkApp3(x_20, x_2, x_3, x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +else +{ +uint8_t x_23; +lean_dec(x_3); +lean_dec(x_2); +x_23 = !lean_is_exclusive(x_13); +if (x_23 == 0) +{ +return x_13; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("of_eq_true", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_1); +x_12 = l_Lean_Meta_Grind_mkEqTrueProof(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3; +x_16 = l_Lean_mkAppB(x_15, x_1, x_14); +lean_ctor_set(x_12, 0, x_16); +return x_12; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_12, 0); +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_12); +x_19 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3; +x_20 = l_Lean_mkAppB(x_19, x_1, x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_12); +if (x_22 == 0) +{ +return x_12; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_12, 0); +x_24 = lean_ctor_get(x_12, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_12); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1___boxed), 10, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_inc(x_1); +x_11 = l_Lean_Meta_instantiateMVarsIfMVarApp(x_1, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1; +x_15 = l_Lean_Expr_cleanupAnnotations(x_12); +x_16 = l_Lean_Expr_isApp(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_15); +x_17 = lean_box(0); +x_18 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Expr_appArg(x_15, lean_box(0)); +x_20 = l_Lean_Expr_appFnCleanup(x_15, lean_box(0)); +x_21 = l_Lean_Expr_isApp(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_19); +x_22 = lean_box(0); +x_23 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = l_Lean_Expr_appArg(x_20, lean_box(0)); +x_25 = l_Lean_Expr_appFnCleanup(x_20, lean_box(0)); +x_26 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3; +x_27 = l_Lean_Expr_isConstOf(x_25, x_26); +if (x_27 == 0) +{ +uint8_t x_28; +lean_dec(x_24); +lean_dec(x_19); +x_28 = l_Lean_Expr_isApp(x_25); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_25); +x_29 = lean_box(0); +x_30 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_30; +} +else +{ +lean_object* x_31; uint8_t x_32; +x_31 = l_Lean_Expr_appFnCleanup(x_25, lean_box(0)); +x_32 = l_Lean_Expr_isApp(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +x_33 = lean_box(0); +x_34 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Expr_appArg(x_31, lean_box(0)); +x_36 = l_Lean_Expr_appFnCleanup(x_31, lean_box(0)); +x_37 = l_Lean_Expr_isApp(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_36); +lean_dec(x_35); +x_38 = lean_box(0); +x_39 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = l_Lean_Expr_appFnCleanup(x_36, lean_box(0)); +x_41 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7; +x_42 = l_Lean_Expr_isConstOf(x_40, x_41); +if (x_42 == 0) +{ +lean_object* x_43; uint8_t x_44; +x_43 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9; +x_44 = l_Lean_Expr_isConstOf(x_40, x_43); +lean_dec(x_40); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_35); +x_45 = lean_box(0); +x_46 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_46; +} +else +{ +lean_object* x_47; +lean_dec(x_1); +x_47 = lean_apply_10(x_14, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_47; +} +} +else +{ +lean_object* x_48; +lean_dec(x_40); +lean_dec(x_1); +x_48 = lean_apply_10(x_14, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_48; +} +} +} +} +} +else +{ +lean_object* x_49; +lean_dec(x_25); +x_49 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2(x_1, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_49; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_introNewHyp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_12 = l_List_reverse___rarg(x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_16 = l_Lean_Meta_Grind_intros(x_3, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_List_appendTR___rarg(x_17, x_2); +x_1 = x_15; +x_2 = x_19; +x_11 = x_18; +goto _start; +} +else +{ +uint8_t x_21; +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) +{ +return x_16; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_16); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; +x_4 = l_List_reverse___rarg(x_3); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_1, 3); +x_11 = lean_ctor_get(x_1, 4); +x_12 = lean_ctor_get(x_1, 5); +x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*19); +x_14 = lean_ctor_get(x_1, 6); +x_15 = lean_ctor_get(x_1, 7); +x_16 = lean_ctor_get(x_1, 8); +x_17 = lean_ctor_get(x_1, 9); +x_18 = lean_ctor_get(x_1, 10); +x_19 = lean_ctor_get(x_1, 11); +x_20 = lean_ctor_get(x_1, 12); +x_21 = lean_ctor_get(x_1, 13); +x_22 = lean_ctor_get(x_1, 14); +x_23 = lean_ctor_get(x_1, 15); +x_24 = lean_ctor_get(x_1, 16); +x_25 = lean_ctor_get(x_1, 17); +x_26 = lean_ctor_get(x_1, 18); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_27 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_27, 0, x_6); +lean_ctor_set(x_27, 1, x_8); +lean_ctor_set(x_27, 2, x_9); +lean_ctor_set(x_27, 3, x_10); +lean_ctor_set(x_27, 4, x_11); +lean_ctor_set(x_27, 5, x_12); +lean_ctor_set(x_27, 6, x_14); +lean_ctor_set(x_27, 7, x_15); +lean_ctor_set(x_27, 8, x_16); +lean_ctor_set(x_27, 9, x_17); +lean_ctor_set(x_27, 10, x_18); +lean_ctor_set(x_27, 11, x_19); +lean_ctor_set(x_27, 12, x_20); +lean_ctor_set(x_27, 13, x_21); +lean_ctor_set(x_27, 14, x_22); +lean_ctor_set(x_27, 15, x_23); +lean_ctor_set(x_27, 16, x_24); +lean_ctor_set(x_27, 17, x_25); +lean_ctor_set(x_27, 18, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*19, x_13); +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_27); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_29 = lean_ctor_get(x_2, 0); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_2); +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_ctor_get(x_1, 2); +x_33 = lean_ctor_get(x_1, 3); +x_34 = lean_ctor_get(x_1, 4); +x_35 = lean_ctor_get(x_1, 5); +x_36 = lean_ctor_get_uint8(x_1, sizeof(void*)*19); +x_37 = lean_ctor_get(x_1, 6); +x_38 = lean_ctor_get(x_1, 7); +x_39 = lean_ctor_get(x_1, 8); +x_40 = lean_ctor_get(x_1, 9); +x_41 = lean_ctor_get(x_1, 10); +x_42 = lean_ctor_get(x_1, 11); +x_43 = lean_ctor_get(x_1, 12); +x_44 = lean_ctor_get(x_1, 13); +x_45 = lean_ctor_get(x_1, 14); +x_46 = lean_ctor_get(x_1, 15); +x_47 = lean_ctor_get(x_1, 16); +x_48 = lean_ctor_get(x_1, 17); +x_49 = lean_ctor_get(x_1, 18); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +x_50 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_50, 0, x_29); +lean_ctor_set(x_50, 1, x_31); +lean_ctor_set(x_50, 2, x_32); +lean_ctor_set(x_50, 3, x_33); +lean_ctor_set(x_50, 4, x_34); +lean_ctor_set(x_50, 5, x_35); +lean_ctor_set(x_50, 6, x_37); +lean_ctor_set(x_50, 7, x_38); +lean_ctor_set(x_50, 8, x_39); +lean_ctor_set(x_50, 9, x_40); +lean_ctor_set(x_50, 10, x_41); +lean_ctor_set(x_50, 11, x_42); +lean_ctor_set(x_50, 12, x_43); +lean_ctor_set(x_50, 13, x_44); +lean_ctor_set(x_50, 14, x_45); +lean_ctor_set(x_50, 15, x_46); +lean_ctor_set(x_50, 16, x_47); +lean_ctor_set(x_50, 17, x_48); +lean_ctor_set(x_50, 18, x_49); +lean_ctor_set_uint8(x_50, sizeof(void*)*19, x_36); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_3); +x_2 = x_30; +x_3 = x_51; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_mk_ref(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_box(0); +x_16 = l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1(x_13, x_2, x_15); +lean_dec(x_13); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_1, x_17); +x_19 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_introNewHyp(x_16, x_15, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_19, 0, x_22); +return x_19; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_19, 0); +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_19); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_23); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +return x_26; +} +} +else +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_19); +if (x_27 == 0) +{ +return x_19; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_19, 0); +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_19); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = l_Lean_Meta_isMatcherApp___at___private_Lean_Meta_Tactic_Grind_Internalize_0__Lean_Meta_Grind_checkAndAddSplitCandidate___spec__1(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_17 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_st_ref_get(x_4, x_19); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_24 = l_Lean_Meta_Grind_cases(x_23, x_18, x_8, x_9, x_10, x_11, x_22); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_Meta_Grind_splitNext___lambda__2(x_2, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +lean_dec(x_4); +return x_27; +} +else +{ +uint8_t x_28; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) +{ +return x_24; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_24, 0); +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_24); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +else +{ +uint8_t x_32; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_32 = !lean_is_exclusive(x_17); +if (x_32 == 0) +{ +return x_17; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_17, 0); +x_34 = lean_ctor_get(x_17, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_17); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_36 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_dec(x_13); +x_37 = lean_st_ref_get(x_4, x_36); +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_41 = l_Lean_Meta_Grind_casesMatch(x_40, x_1, x_8, x_9, x_10, x_11, x_39); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Lean_Meta_Grind_splitNext___lambda__2(x_2, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_43); +lean_dec(x_4); +return x_44; +} +else +{ +uint8_t x_45; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_45 = !lean_is_exclusive(x_41); +if (x_45 == 0) +{ +return x_41; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_41, 0); +x_47 = lean_ctor_get(x_41, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_41); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", generation: ", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_21; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_21 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_box(0); +x_10 = x_24; +x_11 = x_23; +goto block_20; +} +else +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = !lean_is_exclusive(x_22); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_22, 0); +lean_inc(x_27); +x_28 = l_Lean_Meta_Grind_getGeneration(x_27, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_25); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__1; +x_32 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_31, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_unbox(x_33); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_free_object(x_22); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_dec(x_32); +x_36 = lean_box(0); +lean_inc(x_1); +x_37 = l_Lean_Meta_Grind_splitNext___lambda__3(x_27, x_29, x_36, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_35); +lean_dec(x_29); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_10 = x_38; +x_11 = x_39; +goto block_20; +} +else +{ +uint8_t x_40; +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_37); +if (x_40 == 0) +{ +return x_37; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 0); +x_42 = lean_ctor_get(x_37, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_37); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +else +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_32); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_32, 1); +x_46 = lean_ctor_get(x_32, 0); +lean_dec(x_46); +x_47 = l_Lean_Meta_Grind_updateLastTag(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_45); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +lean_inc(x_27); +x_49 = l_Lean_MessageData_ofExpr(x_27); +x_50 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +lean_ctor_set_tag(x_32, 7); +lean_ctor_set(x_32, 1, x_49); +lean_ctor_set(x_32, 0, x_50); +x_51 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__3; +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_32); +lean_ctor_set(x_52, 1, x_51); +lean_inc(x_29); +x_53 = l___private_Init_Data_Repr_0__Nat_reprFast(x_29); +lean_ctor_set_tag(x_22, 3); +lean_ctor_set(x_22, 0, x_53); +x_54 = l_Lean_MessageData_ofFormat(x_22); +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_50); +x_57 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_31, x_56, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +lean_inc(x_1); +x_60 = l_Lean_Meta_Grind_splitNext___lambda__3(x_27, x_29, x_58, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_59); +lean_dec(x_58); +lean_dec(x_29); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_10 = x_61; +x_11 = x_62; +goto block_20; +} +else +{ +uint8_t x_63; +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_60); +if (x_63 == 0) +{ +return x_60; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_60, 0); +x_65 = lean_ctor_get(x_60, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_60); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +else +{ +uint8_t x_67; +lean_free_object(x_32); +lean_dec(x_29); +lean_free_object(x_22); +lean_dec(x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_47); +if (x_67 == 0) +{ +return x_47; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_47, 0); +x_69 = lean_ctor_get(x_47, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_47); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +else +{ +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_32, 1); +lean_inc(x_71); +lean_dec(x_32); +x_72 = l_Lean_Meta_Grind_updateLastTag(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_71); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_73 = lean_ctor_get(x_72, 1); +lean_inc(x_73); +lean_dec(x_72); +lean_inc(x_27); +x_74 = l_Lean_MessageData_ofExpr(x_27); +x_75 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_74); +x_77 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__3; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +lean_inc(x_29); +x_79 = l___private_Init_Data_Repr_0__Nat_reprFast(x_29); +lean_ctor_set_tag(x_22, 3); +lean_ctor_set(x_22, 0, x_79); +x_80 = l_Lean_MessageData_ofFormat(x_22); +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_75); +x_83 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_31, x_82, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_73); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +lean_inc(x_1); +x_86 = l_Lean_Meta_Grind_splitNext___lambda__3(x_27, x_29, x_84, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_85); +lean_dec(x_84); +lean_dec(x_29); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_10 = x_87; +x_11 = x_88; +goto block_20; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_1); +x_89 = lean_ctor_get(x_86, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_86, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_91 = x_86; +} else { + lean_dec_ref(x_86); + x_91 = lean_box(0); +} +if (lean_is_scalar(x_91)) { + x_92 = lean_alloc_ctor(1, 2, 0); +} else { + x_92 = x_91; +} +lean_ctor_set(x_92, 0, x_89); +lean_ctor_set(x_92, 1, x_90); +return x_92; +} +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_29); +lean_free_object(x_22); +lean_dec(x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_93 = lean_ctor_get(x_72, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_72, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_95 = x_72; +} else { + lean_dec_ref(x_72); + x_95 = lean_box(0); +} +if (lean_is_scalar(x_95)) { + x_96 = lean_alloc_ctor(1, 2, 0); +} else { + x_96 = x_95; +} +lean_ctor_set(x_96, 0, x_93); +lean_ctor_set(x_96, 1, x_94); +return x_96; +} +} +} +} +else +{ +uint8_t x_97; +lean_free_object(x_22); +lean_dec(x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_97 = !lean_is_exclusive(x_28); +if (x_97 == 0) +{ +return x_28; +} +else +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_28, 0); +x_99 = lean_ctor_get(x_28, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_28); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; +} +} +} +else +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_22, 0); +lean_inc(x_101); +lean_dec(x_22); +lean_inc(x_101); +x_102 = l_Lean_Meta_Grind_getGeneration(x_101, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_25); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_105 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__1; +x_106 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_105, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_104); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_unbox(x_107); +lean_dec(x_107); +if (x_108 == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_106, 1); +lean_inc(x_109); +lean_dec(x_106); +x_110 = lean_box(0); +lean_inc(x_1); +x_111 = l_Lean_Meta_Grind_splitNext___lambda__3(x_101, x_103, x_110, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_109); +lean_dec(x_103); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_10 = x_112; +x_11 = x_113; +goto block_20; +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_1); +x_114 = lean_ctor_get(x_111, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_111, 1); +lean_inc(x_115); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_116 = x_111; +} else { + lean_dec_ref(x_111); + x_116 = lean_box(0); +} +if (lean_is_scalar(x_116)) { + x_117 = lean_alloc_ctor(1, 2, 0); +} else { + x_117 = x_116; +} +lean_ctor_set(x_117, 0, x_114); +lean_ctor_set(x_117, 1, x_115); +return x_117; +} +} +else +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_106, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_119 = x_106; +} else { + lean_dec_ref(x_106); + x_119 = lean_box(0); +} +x_120 = l_Lean_Meta_Grind_updateLastTag(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_118); +if (lean_obj_tag(x_120) == 0) +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_121 = lean_ctor_get(x_120, 1); +lean_inc(x_121); +lean_dec(x_120); +lean_inc(x_101); +x_122 = l_Lean_MessageData_ofExpr(x_101); +x_123 = l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9; +if (lean_is_scalar(x_119)) { + x_124 = lean_alloc_ctor(7, 2, 0); +} else { + x_124 = x_119; + lean_ctor_set_tag(x_124, 7); +} +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_122); +x_125 = l_Lean_Meta_Grind_splitNext___lambda__4___closed__3; +x_126 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +lean_inc(x_103); +x_127 = l___private_Init_Data_Repr_0__Nat_reprFast(x_103); +x_128 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_128, 0, x_127); +x_129 = l_Lean_MessageData_ofFormat(x_128); +x_130 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_130, 0, x_126); +lean_ctor_set(x_130, 1, x_129); +x_131 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_123); +x_132 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_105, x_131, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_121); +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_132, 1); +lean_inc(x_134); +lean_dec(x_132); +lean_inc(x_1); +x_135 = l_Lean_Meta_Grind_splitNext___lambda__3(x_101, x_103, x_133, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_134); +lean_dec(x_133); +lean_dec(x_103); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_135, 1); +lean_inc(x_137); +lean_dec(x_135); +x_10 = x_136; +x_11 = x_137; +goto block_20; +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +lean_dec(x_1); +x_138 = lean_ctor_get(x_135, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_135, 1); +lean_inc(x_139); +if (lean_is_exclusive(x_135)) { + lean_ctor_release(x_135, 0); + lean_ctor_release(x_135, 1); + x_140 = x_135; +} else { + lean_dec_ref(x_135); + x_140 = lean_box(0); +} +if (lean_is_scalar(x_140)) { + x_141 = lean_alloc_ctor(1, 2, 0); +} else { + x_141 = x_140; +} +lean_ctor_set(x_141, 0, x_138); +lean_ctor_set(x_141, 1, x_139); +return x_141; +} +} +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_119); +lean_dec(x_103); +lean_dec(x_101); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_142 = lean_ctor_get(x_120, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_120, 1); +lean_inc(x_143); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_144 = x_120; +} else { + lean_dec_ref(x_120); + x_144 = lean_box(0); +} +if (lean_is_scalar(x_144)) { + x_145 = lean_alloc_ctor(1, 2, 0); +} else { + x_145 = x_144; +} +lean_ctor_set(x_145, 0, x_142); +lean_ctor_set(x_145, 1, x_143); +return x_145; +} +} +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_dec(x_101); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_146 = lean_ctor_get(x_102, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_102, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_148 = x_102; +} else { + lean_dec_ref(x_102); + x_148 = lean_box(0); +} +if (lean_is_scalar(x_148)) { + x_149 = lean_alloc_ctor(1, 2, 0); +} else { + x_149 = x_148; +} +lean_ctor_set(x_149, 0, x_146); +lean_ctor_set(x_149, 1, x_147); +return x_149; +} +} +} +} +else +{ +uint8_t x_150; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_150 = !lean_is_exclusive(x_21); +if (x_150 == 0) +{ +return x_21; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_21, 0); +x_152 = lean_ctor_get(x_21, 1); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_21); +x_153 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_153, 0, x_151); +lean_ctor_set(x_153, 1, x_152); +return x_153; +} +} +block_20: +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_st_ref_get(x_1, x_11); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_10); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_12); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_10); +lean_ctor_set(x_18, 1, x_16); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_splitNext___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_splitNext___lambda__4), 9, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_splitNext___lambda__1___boxed), 9, 1); +lean_closure_set(x_11, 0, x_1); +x_12 = l_Lean_Meta_Grind_splitNext___closed__1; +x_13 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg), 10, 2); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +x_14 = l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(x_10, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +lean_ctor_set(x_14, 0, x_17); +return x_14; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_14); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) +{ +return x_14; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_14); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_List_mapTR_loop___at_Lean_Meta_Grind_splitNext___spec__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_splitNext___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_splitNext___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_3); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_splitNext___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_splitNext___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +lean_object* initialize_Lean_Meta_Tactic_Grind_Types(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Intro(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_Cases(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Grind_CasesMatch(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Grind_Split(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Meta_Tactic_Grind_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Intro(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_Cases(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Grind_CasesMatch(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1 = _init_l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_CaseSplitStatus_noConfusion___rarg___closed__1); +l_Lean_Meta_Grind_instInhabitedCaseSplitStatus = _init_l_Lean_Meta_Grind_instInhabitedCaseSplitStatus(); +l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1 = _init_l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqCaseSplitStatus___closed__1); +l_Lean_Meta_Grind_instBEqCaseSplitStatus = _init_l_Lean_Meta_Grind_instBEqCaseSplitStatus(); +lean_mark_persistent(l_Lean_Meta_Grind_instBEqCaseSplitStatus); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__1); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__2); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__3); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__4); +l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5 = _init_l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5(); +lean_mark_persistent(l_panic___at___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___spec__1___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__4); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__5___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__4); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__6); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__7); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__8); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___lambda__8___closed__9); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__4); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__6); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__7); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__8); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_checkCaseSplitStatus___closed__9); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_selectNextSplit_x3f___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__4); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__2___closed__5); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__1); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__2); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___lambda__3___closed__3); +l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Split_0__Lean_Meta_Grind_mkCasesMajor___closed__1); +l_Lean_Meta_Grind_splitNext___lambda__4___closed__1 = _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_splitNext___lambda__4___closed__1); +l_Lean_Meta_Grind_splitNext___lambda__4___closed__2 = _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_splitNext___lambda__4___closed__2); +l_Lean_Meta_Grind_splitNext___lambda__4___closed__3 = _init_l_Lean_Meta_Grind_splitNext___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_splitNext___lambda__4___closed__3); +l_Lean_Meta_Grind_splitNext___closed__1 = _init_l_Lean_Meta_Grind_splitNext___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_splitNext___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c index 89336ad917da..1438f7470df4 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Types.c @@ -13,41 +13,60 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Grind_instInhabitedMethods___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedENode___closed__1; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4; +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__1; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instReprENode; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ENode_isCongrRoot___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsInconsistent___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_abstractNestedProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTrueExpr___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__4; static size_t l_Lean_Meta_Grind_instInhabitedGoal___closed__3; +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_canon___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash_goEq___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_abstractNestedProofs___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45; uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethodsRef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -58,104 +77,139 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___lambda__1(lean_object* static lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1; uint64_t lean_uint64_of_nat(lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run(lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46; size_t lean_uint64_to_usize(uint64_t); uint64_t lean_uint64_lor(uint64_t, uint64_t); -uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__5; static lean_object* l_Lean_Meta_Grind_getENode___closed__2; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMainDeclName___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24; uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118_(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getId(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50; uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsInconsistent___closed__1; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7; size_t lean_usize_mul(size_t, size_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_uint64_dec_lt(uint64_t, uint64_t); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5; uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isSameExpr_unsafe__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3; lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getFalseExpr___boxed(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_grind_debug_proofs; +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24; +lean_object* l_Lean_MVarId_getTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachEqc(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Queue_enqueue___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_hasSameRoot(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19; +uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedNewFact___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMainDeclName___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_mkHCongrWithArity___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getFalseExpr(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(lean_object*, lean_object*, lean_object*); static uint64_t l_Lean_Meta_Grind_hasSameType___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrKey___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrTheoremCacheKey___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_grind_debug; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4; lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_AbstractNestedProofs_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isResolvedCaseSplit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instHashablePreInstance___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); @@ -167,15 +221,15 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Queue_empty(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__5; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getFalseExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -183,18 +237,22 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1___boxe LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_isUnaryNode___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__1; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27; +static lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqPreInstance___lambda__1(lean_object*); -static lean_object* l_Lean_Meta_Grind_getENode___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameRoot___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -202,14 +260,22 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__2; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2___boxed(lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_usize_to_uint64(size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25; lean_object* lean_nat_to_int(lean_object*); +lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -218,69 +284,96 @@ static lean_object* l_Lean_Meta_Grind_abstractNestedProofs___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_admit(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHCongrWithArityForConst_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17; +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__4; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31; +lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +lean_object* l_Lean_Expr_appArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTrueExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_quoteNameMk(lean_object*); +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_goEq(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14; lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19; +lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent___boxed(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMainDeclName(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35; static lean_object* l_Lean_Meta_Grind_canon___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20; +lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent_goEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey(lean_object*); LEAN_EXPORT uint64_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_mkHCongrWithArity___spec__5(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39; +static lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isSameExpr___boxed(lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethodsRef___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqCongrTheoremCacheKey(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_canon___closed__1; static lean_object* l_Lean_Meta_Grind_congrHash___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20; +uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqENodeKey___boxed(lean_object*, lean_object*); @@ -288,8 +381,8 @@ lean_object* l_Lean_Meta_mkHCongrWithArity(lean_object*, lean_object*, lean_obje LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent_go___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -297,81 +390,115 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes__ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMainDeclName___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__1; lean_object* l_Lean_Expr_appFn_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyFallback(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Grind_GoalM_run___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19; lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrKey___boxed(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(lean_object*); lean_object* l_Lean_Meta_Grind_Canon_canonImpl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10; +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2(lean_object*, size_t, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFalseElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedENode___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48; lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent_go(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_getENode___closed__3; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrKey(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instReprENode___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_shareCommon___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isResolvedCaseSplit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_GoalM_run_x27___closed__1; +static lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_admit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_ENode_isCongrRoot(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_abstractNestedProofs___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15; LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16; lean_object* lean_grind_mk_eq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFalse(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29; lean_object* lean_grind_mk_heq_proof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_go(lean_object*, lean_object*, uint64_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34; LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Meta_Grind_GoalM_run___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__3; +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1(lean_object*); +lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); +static lean_object* l_Lean_Meta_Grind_congrHash___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_abstractNestedProofs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -382,46 +509,68 @@ static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_mkHCongrWithArity___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods___boxed(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedMethods___closed__1; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isConstructorAppCore_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isSameExpr(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_markAsInconsistent___closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34; static lean_object* l_Lean_Meta_Grind_congrPlaceholderProof___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59; lean_object* l_Lean_indentExpr(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedENode; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__2; +lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqCongrKey(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2; +static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(lean_object*, lean_object*); @@ -432,67 +581,88 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_left(uint64_t, uint64_t); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTrueExpr___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTrueExpr(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isTrueExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_erase_macro_scopes(lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17; +static lean_object* l_Lean_Meta_Grind_congrHash___closed__4; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47; LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42; +lean_object* l_String_intercalate(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22; lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isTrue(lean_object*); uint8_t l_Lean_Expr_quickComp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash_go___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17; +lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(lean_object*); size_t lean_usize_add(size_t, size_t); +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedNewFact; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26; static lean_object* l_Lean_Meta_Grind_getENode___closed__1; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMaxGeneration(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51; +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent_goEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_instInhabitedGoal___closed__4; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__6(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1___boxed(lean_object*); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_118____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); @@ -500,16 +670,26 @@ LEAN_EXPORT uint64_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind static lean_object* l_Lean_Meta_Grind_congrHash___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHCongrWithArity___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3; +extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_KeyedDeclsAttribute_instInhabitedExtensionState___spec__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqPreInstance___lambda__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_canon(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47; static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__5; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11; extern lean_object* l_Lean_ShareCommon_objectFactory; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isFalseExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -517,54 +697,68 @@ lean_object* lean_state_sharecommon(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_mkHCongrWithArity___spec__1(lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__7; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40; +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Meta_TransparencyMode_toUInt64(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrPlaceholderProof; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4; -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25; +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRefPointed; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____; +static lean_object* l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10; static lean_object* l_Lean_Meta_Grind_canon___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run_x27___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_String_toSubstring_x27(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_GoalM_run___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getFalseExpr___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableENodeKey___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT size_t l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey_unsafe__1(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); lean_object* l_Lean_Meta_isLitValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1; +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInterpreted___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_updateLastTag___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__6___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedGoal; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrTheoremCacheKey_unsafe__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getConfig___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Option_repr___at___private_Lean_Meta_Transform_0__Lean_reprTransformStep____x40_Lean_Meta_Transform___hyg_46____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_mkHCongrWithArity___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, uint64_t); +static lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isSameExpr_unsafe__1(lean_object* x_1, lean_object* x_2) { _start: { @@ -1382,13 +1576,15 @@ return x_41; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; x_42 = lean_ctor_get(x_33, 0); x_43 = lean_ctor_get(x_33, 1); x_44 = lean_ctor_get(x_33, 3); x_45 = lean_ctor_get(x_33, 4); x_46 = lean_ctor_get(x_33, 5); x_47 = lean_ctor_get(x_33, 6); +x_48 = lean_ctor_get(x_33, 7); +lean_inc(x_48); lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); @@ -1396,57 +1592,58 @@ lean_inc(x_44); lean_inc(x_43); lean_inc(x_42); lean_dec(x_33); -x_48 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_48, 0, x_42); -lean_ctor_set(x_48, 1, x_43); -lean_ctor_set(x_48, 2, x_30); -lean_ctor_set(x_48, 3, x_44); -lean_ctor_set(x_48, 4, x_45); -lean_ctor_set(x_48, 5, x_46); -lean_ctor_set(x_48, 6, x_47); -x_49 = lean_st_ref_set(x_4, x_48, x_34); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_51 = x_49; +x_49 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_49, 0, x_42); +lean_ctor_set(x_49, 1, x_43); +lean_ctor_set(x_49, 2, x_30); +lean_ctor_set(x_49, 3, x_44); +lean_ctor_set(x_49, 4, x_45); +lean_ctor_set(x_49, 5, x_46); +lean_ctor_set(x_49, 6, x_47); +lean_ctor_set(x_49, 7, x_48); +x_50 = lean_st_ref_set(x_4, x_49, x_34); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_52 = x_50; } else { - lean_dec_ref(x_49); - x_51 = lean_box(0); + lean_dec_ref(x_50); + x_52 = lean_box(0); } -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); } else { - x_52 = x_51; + x_53 = x_52; } -lean_ctor_set(x_52, 0, x_25); -lean_ctor_set(x_52, 1, x_50); -return x_52; +lean_ctor_set(x_53, 0, x_25); +lean_ctor_set(x_53, 1, x_51); +return x_53; } } else { -uint8_t x_53; +uint8_t x_54; lean_dec(x_22); lean_dec(x_18); -x_53 = !lean_is_exclusive(x_24); -if (x_53 == 0) +x_54 = !lean_is_exclusive(x_24); +if (x_54 == 0) { return x_24; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_24, 0); -x_55 = lean_ctor_get(x_24, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_24, 0); +x_56 = lean_ctor_get(x_24, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); lean_dec(x_24); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } } @@ -1509,7 +1706,7 @@ return x_23; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; x_24 = lean_ctor_get(x_11, 0); x_25 = lean_ctor_get(x_11, 1); x_26 = lean_ctor_get(x_11, 2); @@ -1517,6 +1714,8 @@ x_27 = lean_ctor_get(x_11, 3); x_28 = lean_ctor_get(x_11, 4); x_29 = lean_ctor_get(x_11, 5); x_30 = lean_ctor_get(x_11, 6); +x_31 = lean_ctor_get(x_11, 7); +lean_inc(x_31); lean_inc(x_30); lean_inc(x_29); lean_inc(x_28); @@ -1525,40 +1724,41 @@ lean_inc(x_26); lean_inc(x_25); lean_inc(x_24); lean_dec(x_11); -x_31 = l_Lean_ShareCommon_objectFactory; -x_32 = lean_state_sharecommon(x_31, x_25, x_1); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = l_Lean_ShareCommon_objectFactory; +x_33 = lean_state_sharecommon(x_32, x_25, x_1); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_35, 0, x_24); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_26); -lean_ctor_set(x_35, 3, x_27); -lean_ctor_set(x_35, 4, x_28); -lean_ctor_set(x_35, 5, x_29); -lean_ctor_set(x_35, 6, x_30); -x_36 = lean_st_ref_set(x_4, x_35, x_12); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_38 = x_36; +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_36, 0, x_24); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_26); +lean_ctor_set(x_36, 3, x_27); +lean_ctor_set(x_36, 4, x_28); +lean_ctor_set(x_36, 5, x_29); +lean_ctor_set(x_36, 6, x_30); +lean_ctor_set(x_36, 7, x_31); +x_37 = lean_st_ref_set(x_4, x_36, x_12); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_39 = x_37; } else { - lean_dec_ref(x_36); - x_38 = lean_box(0); + lean_dec_ref(x_37); + x_39 = lean_box(0); } -if (lean_is_scalar(x_38)) { - x_39 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(0, 2, 0); } else { - x_39 = x_38; + x_40 = x_39; } -lean_ctor_set(x_39, 0, x_33); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_38); +return x_40; } } } @@ -1679,13 +1879,15 @@ return x_32; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_33 = lean_ctor_get(x_24, 1); x_34 = lean_ctor_get(x_24, 2); x_35 = lean_ctor_get(x_24, 3); x_36 = lean_ctor_get(x_24, 4); x_37 = lean_ctor_get(x_24, 5); x_38 = lean_ctor_get(x_24, 6); +x_39 = lean_ctor_get(x_24, 7); +lean_inc(x_39); lean_inc(x_38); lean_inc(x_37); lean_inc(x_36); @@ -1693,189 +1895,197 @@ lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); lean_dec(x_24); -x_39 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_33); -lean_ctor_set(x_39, 2, x_34); -lean_ctor_set(x_39, 3, x_35); -lean_ctor_set(x_39, 4, x_36); -lean_ctor_set(x_39, 5, x_37); -lean_ctor_set(x_39, 6, x_38); -x_40 = lean_st_ref_set(x_4, x_39, x_25); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_42 = x_40; +x_40 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_33); +lean_ctor_set(x_40, 2, x_34); +lean_ctor_set(x_40, 3, x_35); +lean_ctor_set(x_40, 4, x_36); +lean_ctor_set(x_40, 5, x_37); +lean_ctor_set(x_40, 6, x_38); +lean_ctor_set(x_40, 7, x_39); +x_41 = lean_st_ref_set(x_4, x_40, x_25); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_43 = x_41; } else { - lean_dec_ref(x_40); - x_42 = lean_box(0); + lean_dec_ref(x_41); + x_43 = lean_box(0); } -if (lean_is_scalar(x_42)) { - x_43 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(0, 2, 0); } else { - x_43 = x_42; + x_44 = x_43; } -lean_ctor_set(x_43, 0, x_21); -lean_ctor_set(x_43, 1, x_41); -return x_43; +lean_ctor_set(x_44, 0, x_21); +lean_ctor_set(x_44, 1, x_42); +return x_44; } } else { -uint8_t x_44; -x_44 = !lean_is_exclusive(x_18); -if (x_44 == 0) +uint8_t x_45; +x_45 = !lean_is_exclusive(x_18); +if (x_45 == 0) { return x_18; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_18, 0); -x_46 = lean_ctor_get(x_18, 1); +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_18, 0); +x_47 = lean_ctor_get(x_18, 1); +lean_inc(x_47); lean_inc(x_46); -lean_inc(x_45); lean_dec(x_18); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_48 = lean_ctor_get(x_11, 0); -x_49 = lean_ctor_get(x_11, 1); -x_50 = lean_ctor_get(x_11, 2); -x_51 = lean_ctor_get(x_11, 3); -x_52 = lean_ctor_get(x_11, 4); -x_53 = lean_ctor_get(x_11, 5); -x_54 = lean_ctor_get(x_11, 6); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_49 = lean_ctor_get(x_11, 0); +x_50 = lean_ctor_get(x_11, 1); +x_51 = lean_ctor_get(x_11, 2); +x_52 = lean_ctor_get(x_11, 3); +x_53 = lean_ctor_get(x_11, 4); +x_54 = lean_ctor_get(x_11, 5); +x_55 = lean_ctor_get(x_11, 6); +x_56 = lean_ctor_get(x_11, 7); +lean_inc(x_56); +lean_inc(x_55); lean_inc(x_54); lean_inc(x_53); lean_inc(x_52); lean_inc(x_51); lean_inc(x_50); lean_inc(x_49); -lean_inc(x_48); lean_dec(x_11); -x_55 = l_Lean_Meta_Grind_canon___closed__3; -x_56 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_49); -lean_ctor_set(x_56, 2, x_50); -lean_ctor_set(x_56, 3, x_51); -lean_ctor_set(x_56, 4, x_52); -lean_ctor_set(x_56, 5, x_53); -lean_ctor_set(x_56, 6, x_54); -x_57 = lean_st_ref_set(x_4, x_56, x_12); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_48, x_5, x_6, x_7, x_8, x_58); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_60 = lean_ctor_get(x_59, 0); +x_57 = l_Lean_Meta_Grind_canon___closed__3; +x_58 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_50); +lean_ctor_set(x_58, 2, x_51); +lean_ctor_set(x_58, 3, x_52); +lean_ctor_set(x_58, 4, x_53); +lean_ctor_set(x_58, 5, x_54); +lean_ctor_set(x_58, 6, x_55); +lean_ctor_set(x_58, 7, x_56); +x_59 = lean_st_ref_set(x_4, x_58, x_12); +x_60 = lean_ctor_get(x_59, 1); lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); lean_dec(x_59); -x_62 = lean_ctor_get(x_60, 0); +x_61 = l_Lean_Meta_Grind_Canon_canonImpl(x_1, x_49, x_5, x_6, x_7, x_8, x_60); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_62 = lean_ctor_get(x_61, 0); lean_inc(x_62); -x_63 = lean_ctor_get(x_60, 1); +x_63 = lean_ctor_get(x_61, 1); lean_inc(x_63); -lean_dec(x_60); -x_64 = lean_st_ref_take(x_4, x_61); -x_65 = lean_ctor_get(x_64, 0); +lean_dec(x_61); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_62, 1); lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = lean_ctor_get(x_65, 1); +lean_dec(x_62); +x_66 = lean_st_ref_take(x_4, x_63); +x_67 = lean_ctor_get(x_66, 0); lean_inc(x_67); -x_68 = lean_ctor_get(x_65, 2); +x_68 = lean_ctor_get(x_66, 1); lean_inc(x_68); -x_69 = lean_ctor_get(x_65, 3); +lean_dec(x_66); +x_69 = lean_ctor_get(x_67, 1); lean_inc(x_69); -x_70 = lean_ctor_get(x_65, 4); +x_70 = lean_ctor_get(x_67, 2); lean_inc(x_70); -x_71 = lean_ctor_get(x_65, 5); +x_71 = lean_ctor_get(x_67, 3); lean_inc(x_71); -x_72 = lean_ctor_get(x_65, 6); +x_72 = lean_ctor_get(x_67, 4); lean_inc(x_72); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - lean_ctor_release(x_65, 2); - lean_ctor_release(x_65, 3); - lean_ctor_release(x_65, 4); - lean_ctor_release(x_65, 5); - lean_ctor_release(x_65, 6); - x_73 = x_65; +x_73 = lean_ctor_get(x_67, 5); +lean_inc(x_73); +x_74 = lean_ctor_get(x_67, 6); +lean_inc(x_74); +x_75 = lean_ctor_get(x_67, 7); +lean_inc(x_75); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + lean_ctor_release(x_67, 2); + lean_ctor_release(x_67, 3); + lean_ctor_release(x_67, 4); + lean_ctor_release(x_67, 5); + lean_ctor_release(x_67, 6); + lean_ctor_release(x_67, 7); + x_76 = x_67; } else { - lean_dec_ref(x_65); - x_73 = lean_box(0); + lean_dec_ref(x_67); + x_76 = lean_box(0); } -if (lean_is_scalar(x_73)) { - x_74 = lean_alloc_ctor(0, 7, 0); +if (lean_is_scalar(x_76)) { + x_77 = lean_alloc_ctor(0, 8, 0); } else { - x_74 = x_73; -} -lean_ctor_set(x_74, 0, x_63); -lean_ctor_set(x_74, 1, x_67); -lean_ctor_set(x_74, 2, x_68); -lean_ctor_set(x_74, 3, x_69); -lean_ctor_set(x_74, 4, x_70); -lean_ctor_set(x_74, 5, x_71); -lean_ctor_set(x_74, 6, x_72); -x_75 = lean_st_ref_set(x_4, x_74, x_66); -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_77 = x_75; + x_77 = x_76; +} +lean_ctor_set(x_77, 0, x_65); +lean_ctor_set(x_77, 1, x_69); +lean_ctor_set(x_77, 2, x_70); +lean_ctor_set(x_77, 3, x_71); +lean_ctor_set(x_77, 4, x_72); +lean_ctor_set(x_77, 5, x_73); +lean_ctor_set(x_77, 6, x_74); +lean_ctor_set(x_77, 7, x_75); +x_78 = lean_st_ref_set(x_4, x_77, x_68); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_80 = x_78; } else { - lean_dec_ref(x_75); - x_77 = lean_box(0); + lean_dec_ref(x_78); + x_80 = lean_box(0); } -if (lean_is_scalar(x_77)) { - x_78 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 2, 0); } else { - x_78 = x_77; + x_81 = x_80; } -lean_ctor_set(x_78, 0, x_62); -lean_ctor_set(x_78, 1, x_76); -return x_78; +lean_ctor_set(x_81, 0, x_64); +lean_ctor_set(x_81, 1, x_79); +return x_81; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_79 = lean_ctor_get(x_59, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_59, 1); -lean_inc(x_80); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_81 = x_59; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_82 = lean_ctor_get(x_61, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_61, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_84 = x_61; } else { - lean_dec_ref(x_59); - x_81 = lean_box(0); + lean_dec_ref(x_61); + x_84 = lean_box(0); } -if (lean_is_scalar(x_81)) { - x_82 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); } else { - x_82 = x_81; + x_85 = x_84; } -lean_ctor_set(x_82, 0, x_79); -lean_ctor_set(x_82, 1, x_80); -return x_82; +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; } } } @@ -2965,7 +3175,7 @@ return x_26; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; x_27 = lean_ctor_get(x_17, 0); x_28 = lean_ctor_get(x_17, 1); x_29 = lean_ctor_get(x_17, 2); @@ -2973,6 +3183,8 @@ x_30 = lean_ctor_get(x_17, 3); x_31 = lean_ctor_get(x_17, 4); x_32 = lean_ctor_get(x_17, 5); x_33 = lean_ctor_get(x_17, 6); +x_34 = lean_ctor_get(x_17, 7); +lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); @@ -2982,57 +3194,58 @@ lean_inc(x_28); lean_inc(x_27); lean_dec(x_17); lean_inc(x_14); -x_34 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_mkHCongrWithArity___spec__1(x_30, x_3, x_14); -x_35 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_35, 0, x_27); -lean_ctor_set(x_35, 1, x_28); -lean_ctor_set(x_35, 2, x_29); -lean_ctor_set(x_35, 3, x_34); -lean_ctor_set(x_35, 4, x_31); -lean_ctor_set(x_35, 5, x_32); -lean_ctor_set(x_35, 6, x_33); -x_36 = lean_st_ref_set(x_7, x_35, x_18); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_38 = x_36; +x_35 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_mkHCongrWithArity___spec__1(x_30, x_3, x_14); +x_36 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_36, 0, x_27); +lean_ctor_set(x_36, 1, x_28); +lean_ctor_set(x_36, 2, x_29); +lean_ctor_set(x_36, 3, x_35); +lean_ctor_set(x_36, 4, x_31); +lean_ctor_set(x_36, 5, x_32); +lean_ctor_set(x_36, 6, x_33); +lean_ctor_set(x_36, 7, x_34); +x_37 = lean_st_ref_set(x_7, x_36, x_18); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_39 = x_37; } else { - lean_dec_ref(x_36); - x_38 = lean_box(0); + lean_dec_ref(x_37); + x_39 = lean_box(0); } -if (lean_is_scalar(x_38)) { - x_39 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(0, 2, 0); } else { - x_39 = x_38; + x_40 = x_39; } -lean_ctor_set(x_39, 0, x_14); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_ctor_set(x_40, 0, x_14); +lean_ctor_set(x_40, 1, x_38); +return x_40; } } else { -uint8_t x_40; +uint8_t x_41; lean_dec(x_3); -x_40 = !lean_is_exclusive(x_13); -if (x_40 == 0) +x_41 = !lean_is_exclusive(x_13); +if (x_41 == 0) { return x_13; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_13, 0); -x_42 = lean_ctor_get(x_13, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_13, 0); +x_43 = lean_ctor_get(x_13, 1); +lean_inc(x_43); lean_inc(x_42); -lean_inc(x_41); lean_dec(x_13); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } @@ -3070,7 +3283,7 @@ return x_19; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3105,41 +3318,44 @@ x_31 = lean_ctor_get(x_23, 5); lean_inc(x_31); x_32 = lean_ctor_get(x_23, 6); lean_inc(x_32); +x_33 = lean_ctor_get(x_23, 7); +lean_inc(x_33); lean_dec(x_23); -x_33 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_33, 0, x_25); -lean_ctor_set(x_33, 1, x_26); -lean_ctor_set(x_33, 2, x_27); -lean_ctor_set(x_33, 3, x_29); -lean_ctor_set(x_33, 4, x_30); -lean_ctor_set(x_33, 5, x_31); -lean_ctor_set(x_33, 6, x_32); -x_34 = lean_st_ref_set(x_7, x_33, x_24); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +x_34 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_34, 0, x_25); +lean_ctor_set(x_34, 1, x_26); +lean_ctor_set(x_34, 2, x_27); +lean_ctor_set(x_34, 3, x_29); +lean_ctor_set(x_34, 4, x_30); +lean_ctor_set(x_34, 5, x_31); +lean_ctor_set(x_34, 6, x_32); +lean_ctor_set(x_34, 7, x_33); +x_35 = lean_st_ref_set(x_7, x_34, x_24); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -lean_ctor_set(x_34, 0, x_21); -return x_34; +lean_object* x_37; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +lean_ctor_set(x_35, 0, x_21); +return x_35; } else { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); -lean_dec(x_34); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_21); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_dec(x_35); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_21); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } } else { -uint8_t x_39; +uint8_t x_40; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -3147,32 +3363,32 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_39 = !lean_is_exclusive(x_15); -if (x_39 == 0) +x_40 = !lean_is_exclusive(x_15); +if (x_40 == 0) { return x_15; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_15, 0); -x_41 = lean_ctor_get(x_15, 1); +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_15, 0); +x_42 = lean_ctor_get(x_15, 1); +lean_inc(x_42); lean_inc(x_41); -lean_inc(x_40); lean_dec(x_15); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } } else { -lean_object* x_43; lean_object* x_44; -x_43 = lean_box(0); -x_44 = l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1(x_1, x_2, x_3, x_43, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_44; +lean_object* x_44; lean_object* x_45; +x_44 = lean_box(0); +x_45 = l_Lean_Meta_Grind_mkHCongrWithArity___lambda__1(x_1, x_2, x_3, x_44, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_45; } } } @@ -3395,7 +3611,7 @@ x_1 = l_Lean_Meta_Grind_instInhabitedENode___closed__2; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1() { _start: { lean_object* x_1; @@ -3403,29 +3619,29 @@ x_1 = lean_mk_string_unchecked("self", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4() { _start: { lean_object* x_1; @@ -3433,29 +3649,29 @@ x_1 = lean_mk_string_unchecked(" := ", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -3464,7 +3680,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8() { _start: { lean_object* x_1; @@ -3472,17 +3688,17 @@ x_1 = lean_mk_string_unchecked(",", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10() { _start: { lean_object* x_1; @@ -3490,17 +3706,17 @@ x_1 = lean_mk_string_unchecked("next", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12() { _start: { lean_object* x_1; @@ -3508,44 +3724,44 @@ x_1 = lean_mk_string_unchecked("root", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("cgRoot", 6, 6); +x_1 = lean_mk_string_unchecked("congr", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(10u); +x_1 = lean_unsigned_to_nat(9u); x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17() { _start: { lean_object* x_1; @@ -3553,17 +3769,17 @@ x_1 = lean_mk_string_unchecked("target\?", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -3572,7 +3788,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20() { _start: { lean_object* x_1; @@ -3580,17 +3796,26 @@ x_1 = lean_mk_string_unchecked("proof\?", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23() { _start: { lean_object* x_1; @@ -3598,17 +3823,17 @@ x_1 = lean_mk_string_unchecked("flipped", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25() { _start: { lean_object* x_1; @@ -3616,17 +3841,17 @@ x_1 = lean_mk_string_unchecked("size", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27() { _start: { lean_object* x_1; @@ -3634,17 +3859,17 @@ x_1 = lean_mk_string_unchecked("interpreted", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -3653,7 +3878,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30() { _start: { lean_object* x_1; @@ -3661,17 +3886,17 @@ x_1 = lean_mk_string_unchecked("ctor", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32() { _start: { lean_object* x_1; @@ -3679,17 +3904,17 @@ x_1 = lean_mk_string_unchecked("hasLambdas", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -3698,7 +3923,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35() { _start: { lean_object* x_1; @@ -3706,17 +3931,17 @@ x_1 = lean_mk_string_unchecked("heqProofs", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -3725,7 +3950,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38() { _start: { lean_object* x_1; @@ -3733,17 +3958,17 @@ x_1 = lean_mk_string_unchecked("idx", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -3752,7 +3977,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41() { _start: { lean_object* x_1; @@ -3760,17 +3985,17 @@ x_1 = lean_mk_string_unchecked("generation", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43() { _start: { lean_object* x_1; @@ -3778,17 +4003,17 @@ x_1 = lean_mk_string_unchecked("mt", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -3797,7 +4022,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46() { _start: { lean_object* x_1; @@ -3805,35 +4030,35 @@ x_1 = lean_mk_string_unchecked("{ ", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50() { _start: { lean_object* x_1; @@ -3841,17 +4066,17 @@ x_1 = lean_mk_string_unchecked(" }", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52() { _start: { lean_object* x_1; @@ -3859,17 +4084,17 @@ x_1 = lean_mk_string_unchecked("false", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54() { _start: { lean_object* x_1; @@ -3877,25 +4102,25 @@ x_1 = lean_mk_string_unchecked("true", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54() { +static lean_object* _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53; +x_1 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560_(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_unsigned_to_nat(0u); x_5 = l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(x_3, x_4); -x_6 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7; +x_6 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -3903,11 +4128,11 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6; +x_10 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); -x_12 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9; +x_12 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9; x_13 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -3915,11 +4140,11 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11; +x_16 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5; +x_18 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5; x_19 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); @@ -3941,7 +4166,7 @@ lean_ctor_set(x_25, 1, x_12); x_26 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_14); -x_27 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13; +x_27 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13; x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -3966,7 +4191,7 @@ lean_ctor_set(x_35, 1, x_12); x_36 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_14); -x_37 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15; +x_37 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15; x_38 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_37); @@ -3976,7 +4201,7 @@ lean_ctor_set(x_39, 1, x_18); x_40 = lean_ctor_get(x_1, 3); lean_inc(x_40); x_41 = l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3036_(x_40, x_4); -x_42 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16; +x_42 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16; x_43 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_43, 0, x_42); lean_ctor_set(x_43, 1, x_41); @@ -3992,7 +4217,7 @@ lean_ctor_set(x_46, 1, x_12); x_47 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_47, 0, x_46); lean_ctor_set(x_47, 1, x_14); -x_48 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18; +x_48 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18; x_49 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); @@ -4002,7 +4227,7 @@ lean_ctor_set(x_50, 1, x_18); x_51 = lean_ctor_get(x_1, 4); lean_inc(x_51); x_52 = l_Option_repr___at___private_Lean_Meta_Transform_0__Lean_reprTransformStep____x40_Lean_Meta_Transform___hyg_46____spec__1(x_51, x_4); -x_53 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19; +x_53 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19; x_54 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_54, 0, x_53); lean_ctor_set(x_54, 1, x_52); @@ -4018,7 +4243,7 @@ lean_ctor_set(x_57, 1, x_12); x_58 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_58, 0, x_57); lean_ctor_set(x_58, 1, x_14); -x_59 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21; +x_59 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21; x_60 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_60, 0, x_58); lean_ctor_set(x_60, 1, x_59); @@ -4028,345 +4253,346 @@ lean_ctor_set(x_61, 1, x_18); x_62 = lean_ctor_get(x_1, 5); lean_inc(x_62); x_63 = l_Option_repr___at___private_Lean_Meta_Transform_0__Lean_reprTransformStep____x40_Lean_Meta_Transform___hyg_46____spec__1(x_62, x_4); -x_64 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_64, 0, x_42); -lean_ctor_set(x_64, 1, x_63); -x_65 = lean_alloc_ctor(6, 1, 1); +x_64 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22; +x_65 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_65, 0, x_64); -lean_ctor_set_uint8(x_65, sizeof(void*)*1, x_8); -x_66 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_66, 0, x_61); -lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_65, 1, x_63); +x_66 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set_uint8(x_66, sizeof(void*)*1, x_8); x_67 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_12); +lean_ctor_set(x_67, 0, x_61); +lean_ctor_set(x_67, 1, x_66); x_68 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_14); -x_69 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23; -x_70 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); +lean_ctor_set(x_68, 1, x_12); +x_69 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_14); +x_70 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24; x_71 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_18); -x_72 = lean_ctor_get_uint8(x_1, sizeof(void*)*10); -x_73 = lean_ctor_get(x_1, 6); -lean_inc(x_73); -x_74 = l___private_Init_Data_Repr_0__Nat_reprFast(x_73); -x_75 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_75, 0, x_74); -x_76 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_76, 0, x_6); -lean_ctor_set(x_76, 1, x_75); -x_77 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set_uint8(x_77, sizeof(void*)*1, x_8); -x_78 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 1); -x_79 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 2); -x_80 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 3); -x_81 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 4); -x_82 = lean_ctor_get(x_1, 7); -lean_inc(x_82); -x_83 = l___private_Init_Data_Repr_0__Nat_reprFast(x_82); -x_84 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_84, 0, x_83); -x_85 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39; -x_86 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_84); -x_87 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +x_72 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_18); +x_73 = lean_ctor_get_uint8(x_1, sizeof(void*)*10); +x_74 = lean_ctor_get(x_1, 6); +lean_inc(x_74); +x_75 = l___private_Init_Data_Repr_0__Nat_reprFast(x_74); +x_76 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_77, 0, x_6); +lean_ctor_set(x_77, 1, x_76); +x_78 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set_uint8(x_78, sizeof(void*)*1, x_8); +x_79 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 1); +x_80 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 2); +x_81 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 3); +x_82 = lean_ctor_get_uint8(x_1, sizeof(void*)*10 + 4); +x_83 = lean_ctor_get(x_1, 7); +lean_inc(x_83); +x_84 = l___private_Init_Data_Repr_0__Nat_reprFast(x_83); +x_85 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_86 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40; +x_87 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_87, 0, x_86); -lean_ctor_set_uint8(x_87, sizeof(void*)*1, x_8); -x_88 = lean_ctor_get(x_1, 8); -lean_inc(x_88); -x_89 = l___private_Init_Data_Repr_0__Nat_reprFast(x_88); -x_90 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_90, 0, x_89); -x_91 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33; -x_92 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_90); -x_93 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_87, 1, x_85); +x_88 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set_uint8(x_88, sizeof(void*)*1, x_8); +x_89 = lean_ctor_get(x_1, 8); +lean_inc(x_89); +x_90 = l___private_Init_Data_Repr_0__Nat_reprFast(x_89); +x_91 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_91, 0, x_90); +x_92 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34; +x_93 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_93, 0, x_92); -lean_ctor_set_uint8(x_93, sizeof(void*)*1, x_8); -x_94 = lean_ctor_get(x_1, 9); -lean_inc(x_94); +lean_ctor_set(x_93, 1, x_91); +x_94 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set_uint8(x_94, sizeof(void*)*1, x_8); +x_95 = lean_ctor_get(x_1, 9); +lean_inc(x_95); lean_dec(x_1); -x_95 = l___private_Init_Data_Repr_0__Nat_reprFast(x_94); -x_96 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_96, 0, x_95); -x_97 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44; -x_98 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_96); -x_99 = lean_alloc_ctor(6, 1, 1); +x_96 = l___private_Init_Data_Repr_0__Nat_reprFast(x_95); +x_97 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_98 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45; +x_99 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_99, 0, x_98); -lean_ctor_set_uint8(x_99, sizeof(void*)*1, x_8); -if (x_72 == 0) +lean_ctor_set(x_99, 1, x_97); +x_100 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set_uint8(x_100, sizeof(void*)*1, x_8); +if (x_73 == 0) { -lean_object* x_186; -x_186 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_100 = x_186; -goto block_185; +lean_object* x_187; +x_187 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_101 = x_187; +goto block_186; } else { -lean_object* x_187; -x_187 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_100 = x_187; -goto block_185; -} -block_185: -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_101 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_101, 0, x_53); -lean_ctor_set(x_101, 1, x_100); -x_102 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set_uint8(x_102, sizeof(void*)*1, x_8); -x_103 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_103, 0, x_71); -lean_ctor_set(x_103, 1, x_102); +lean_object* x_188; +x_188 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_101 = x_188; +goto block_186; +} +block_186: +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_102 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_102, 0, x_53); +lean_ctor_set(x_102, 1, x_101); +x_103 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set_uint8(x_103, sizeof(void*)*1, x_8); x_104 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_12); +lean_ctor_set(x_104, 0, x_72); +lean_ctor_set(x_104, 1, x_103); x_105 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_14); -x_106 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25; -x_107 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); +lean_ctor_set(x_105, 1, x_12); +x_106 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_14); +x_107 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26; x_108 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_108, 0, x_107); -lean_ctor_set(x_108, 1, x_18); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); x_109 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_77); +lean_ctor_set(x_109, 1, x_18); x_110 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_12); +lean_ctor_set(x_110, 1, x_78); x_111 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_14); -x_112 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27; -x_113 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_111, 1, x_12); +x_112 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_14); +x_113 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28; x_114 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_18); -if (x_78 == 0) +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +x_115 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_18); +if (x_79 == 0) { -lean_object* x_183; -x_183 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_115 = x_183; -goto block_182; +lean_object* x_184; +x_184 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_116 = x_184; +goto block_183; } else { -lean_object* x_184; -x_184 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_115 = x_184; -goto block_182; +lean_object* x_185; +x_185 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_116 = x_185; +goto block_183; } -block_182: +block_183: { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_116 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28; -x_117 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_115); -x_118 = lean_alloc_ctor(6, 1, 1); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_117 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29; +x_118 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_118, 0, x_117); -lean_ctor_set_uint8(x_118, sizeof(void*)*1, x_8); -x_119 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_119, 0, x_114); -lean_ctor_set(x_119, 1, x_118); +lean_ctor_set(x_118, 1, x_116); +x_119 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set_uint8(x_119, sizeof(void*)*1, x_8); x_120 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_12); +lean_ctor_set(x_120, 0, x_115); +lean_ctor_set(x_120, 1, x_119); x_121 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_121, 0, x_120); -lean_ctor_set(x_121, 1, x_14); -x_122 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30; -x_123 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_121, 1, x_12); +x_122 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_14); +x_123 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31; x_124 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_18); -if (x_79 == 0) +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +x_125 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_18); +if (x_80 == 0) { -lean_object* x_180; -x_180 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_125 = x_180; -goto block_179; +lean_object* x_181; +x_181 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_126 = x_181; +goto block_180; } else { -lean_object* x_181; -x_181 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_125 = x_181; -goto block_179; -} -block_179: -{ -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_126 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_126, 0, x_6); -lean_ctor_set(x_126, 1, x_125); -x_127 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_127, 0, x_126); -lean_ctor_set_uint8(x_127, sizeof(void*)*1, x_8); -x_128 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_128, 0, x_124); -lean_ctor_set(x_128, 1, x_127); +lean_object* x_182; +x_182 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_126 = x_182; +goto block_180; +} +block_180: +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_127 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_127, 0, x_6); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set_uint8(x_128, sizeof(void*)*1, x_8); x_129 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_12); +lean_ctor_set(x_129, 0, x_125); +lean_ctor_set(x_129, 1, x_128); x_130 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_130, 0, x_129); -lean_ctor_set(x_130, 1, x_14); -x_131 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32; -x_132 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_130, 1, x_12); +x_131 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_14); +x_132 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33; x_133 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_18); -if (x_80 == 0) +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +x_134 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_18); +if (x_81 == 0) { -lean_object* x_177; -x_177 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_134 = x_177; -goto block_176; +lean_object* x_178; +x_178 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_135 = x_178; +goto block_177; } else { -lean_object* x_178; -x_178 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_134 = x_178; -goto block_176; -} -block_176: -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_135 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_135, 0, x_91); -lean_ctor_set(x_135, 1, x_134); -x_136 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set_uint8(x_136, sizeof(void*)*1, x_8); -x_137 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_137, 0, x_133); -lean_ctor_set(x_137, 1, x_136); +lean_object* x_179; +x_179 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_135 = x_179; +goto block_177; +} +block_177: +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_136 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_136, 0, x_92); +lean_ctor_set(x_136, 1, x_135); +x_137 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set_uint8(x_137, sizeof(void*)*1, x_8); x_138 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_12); +lean_ctor_set(x_138, 0, x_134); +lean_ctor_set(x_138, 1, x_137); x_139 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_139, 0, x_138); -lean_ctor_set(x_139, 1, x_14); -x_140 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35; -x_141 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set(x_141, 1, x_140); +lean_ctor_set(x_139, 1, x_12); +x_140 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_14); +x_141 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36; x_142 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_18); -if (x_81 == 0) +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +x_143 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_18); +if (x_82 == 0) { -lean_object* x_174; -x_174 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52; -x_143 = x_174; -goto block_173; +lean_object* x_175; +x_175 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53; +x_144 = x_175; +goto block_174; } else { -lean_object* x_175; -x_175 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54; -x_143 = x_175; -goto block_173; -} -block_173: -{ -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_144 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36; -x_145 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_145, 0, x_144); -lean_ctor_set(x_145, 1, x_143); -x_146 = lean_alloc_ctor(6, 1, 1); +lean_object* x_176; +x_176 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55; +x_144 = x_176; +goto block_174; +} +block_174: +{ +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_145 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37; +x_146 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_146, 0, x_145); -lean_ctor_set_uint8(x_146, sizeof(void*)*1, x_8); -x_147 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_147, 0, x_142); -lean_ctor_set(x_147, 1, x_146); +lean_ctor_set(x_146, 1, x_144); +x_147 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set_uint8(x_147, sizeof(void*)*1, x_8); x_148 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_12); +lean_ctor_set(x_148, 0, x_143); +lean_ctor_set(x_148, 1, x_147); x_149 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_14); -x_150 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38; -x_151 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); +lean_ctor_set(x_149, 1, x_12); +x_150 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_14); +x_151 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39; x_152 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_152, 0, x_151); -lean_ctor_set(x_152, 1, x_18); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); x_153 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_87); +lean_ctor_set(x_153, 1, x_18); x_154 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_12); +lean_ctor_set(x_154, 1, x_88); x_155 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_14); -x_156 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41; -x_157 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_157, 0, x_155); -lean_ctor_set(x_157, 1, x_156); +lean_ctor_set(x_155, 1, x_12); +x_156 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_14); +x_157 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42; x_158 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_18); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); x_159 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_159, 0, x_158); -lean_ctor_set(x_159, 1, x_93); +lean_ctor_set(x_159, 1, x_18); x_160 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_12); +lean_ctor_set(x_160, 1, x_94); x_161 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_14); -x_162 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43; -x_163 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_163, 0, x_161); -lean_ctor_set(x_163, 1, x_162); +lean_ctor_set(x_161, 1, x_12); +x_162 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_162, 0, x_161); +lean_ctor_set(x_162, 1, x_14); +x_163 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44; x_164 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_18); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); x_165 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_99); -x_166 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48; -x_167 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_167, 0, x_166); -lean_ctor_set(x_167, 1, x_165); -x_168 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50; -x_169 = lean_alloc_ctor(5, 2, 0); -lean_ctor_set(x_169, 0, x_167); -lean_ctor_set(x_169, 1, x_168); -x_170 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47; -x_171 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_171, 0, x_170); -lean_ctor_set(x_171, 1, x_169); -x_172 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_165, 1, x_18); +x_166 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_100); +x_167 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49; +x_168 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_166); +x_169 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51; +x_170 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_169); +x_171 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48; +x_172 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_172, 0, x_171); -lean_ctor_set_uint8(x_172, sizeof(void*)*1, x_8); -return x_172; +lean_ctor_set(x_172, 1, x_170); +x_173 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set_uint8(x_173, sizeof(void*)*1, x_8); +return x_173; } } } @@ -4374,11 +4600,11 @@ return x_172; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4387,7 +4613,7 @@ static lean_object* _init_l_Lean_Meta_Grind_instReprENode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____boxed), 2, 0); return x_1; } } @@ -4399,6 +4625,26 @@ x_1 = l_Lean_Meta_Grind_instReprENode___closed__1; return x_1; } } +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_ENode_isCongrRoot(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_1, 3); +x_4 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_ENode_isCongrRoot___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(lean_object* x_1) { _start: { @@ -4733,7 +4979,7 @@ x_4 = lean_box_uint64(x_3); return x_4; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_hasSameRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -4792,17 +5038,50 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; -x_4 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); +x_4 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); x_5 = lean_box(x_4); return x_5; } } +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_goEq(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; uint64_t x_5; uint8_t x_6; +lean_inc(x_1); +x_4 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(x_1, x_2); +x_5 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(x_1, x_3); +x_6 = lean_uint64_dec_lt(x_5, x_4); +if (x_6 == 0) +{ +uint64_t x_7; +x_7 = lean_uint64_mix_hash(x_4, x_5); +return x_7; +} +else +{ +uint64_t x_8; +x_8 = lean_uint64_mix_hash(x_5, x_4); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash_goEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; lean_object* x_5; +x_4 = l_Lean_Meta_Grind_congrHash_goEq(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box_uint64(x_4); +return x_5; +} +} LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash_go(lean_object* x_1, lean_object* x_2, uint64_t x_3) { _start: { @@ -4858,34 +5137,117 @@ x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_congrHash___closed__3() { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = l_Lean_Meta_Grind_congrHash___closed__2; -x_4 = lean_unsigned_to_nat(2u); -x_5 = l_Lean_Expr_isAppOfArity(x_2, x_3, x_4); -if (x_5 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Eq", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_congrHash___closed__4() { +_start: { -uint64_t x_6; uint64_t x_7; -x_6 = 17; -x_7 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_6); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_congrHash___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_congrHash(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +lean_inc(x_2); +x_3 = l_Lean_Expr_cleanupAnnotations(x_2); +x_4 = l_Lean_Expr_isApp(x_3); +if (x_4 == 0) +{ +uint64_t x_5; uint64_t x_6; +lean_dec(x_3); +x_5 = 17; +x_6 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_5); +lean_dec(x_2); +return x_6; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint64_t x_14; -x_8 = lean_unsigned_to_nat(0u); -x_9 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_8); -x_10 = lean_nat_sub(x_9, x_8); -lean_dec(x_9); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_sub(x_10, x_11); -lean_dec(x_10); -x_13 = l_Lean_Expr_getRevArg_x21(x_2, x_12); -x_14 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(x_1, x_13); +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_Expr_appArg(x_3, lean_box(0)); +x_8 = l_Lean_Expr_appFnCleanup(x_3, lean_box(0)); +x_9 = l_Lean_Expr_isApp(x_8); +if (x_9 == 0) +{ +uint64_t x_10; uint64_t x_11; +lean_dec(x_8); +lean_dec(x_7); +x_10 = 17; +x_11 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_10); +lean_dec(x_2); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = l_Lean_Expr_appArg(x_8, lean_box(0)); +x_13 = l_Lean_Expr_appFnCleanup(x_8, lean_box(0)); +x_14 = l_Lean_Meta_Grind_congrHash___closed__2; +x_15 = l_Lean_Expr_isConstOf(x_13, x_14); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = l_Lean_Expr_isApp(x_13); +if (x_16 == 0) +{ +uint64_t x_17; uint64_t x_18; lean_dec(x_13); -return x_14; +lean_dec(x_12); +lean_dec(x_7); +x_17 = 17; +x_18 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_17); +lean_dec(x_2); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Expr_appFnCleanup(x_13, lean_box(0)); +x_20 = l_Lean_Meta_Grind_congrHash___closed__4; +x_21 = l_Lean_Expr_isConstOf(x_19, x_20); +lean_dec(x_19); +if (x_21 == 0) +{ +uint64_t x_22; uint64_t x_23; +lean_dec(x_12); +lean_dec(x_7); +x_22 = 17; +x_23 = l_Lean_Meta_Grind_congrHash_go(x_1, x_2, x_22); +lean_dec(x_2); +return x_23; +} +else +{ +uint64_t x_24; +lean_dec(x_2); +x_24 = l_Lean_Meta_Grind_congrHash_goEq(x_1, x_12, x_7); +lean_dec(x_7); +lean_dec(x_12); +return x_24; +} +} +} +else +{ +uint64_t x_25; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_2); +x_25 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot(x_1, x_12); +lean_dec(x_12); +return x_25; +} +} } } } @@ -4894,11 +5256,82 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Grind_congrHash___boxed(lean_object* x_1, l { uint64_t x_3; lean_object* x_4; x_3 = l_Lean_Meta_Grind_congrHash(x_1, x_2); -lean_dec(x_2); x_4 = lean_box_uint64(x_3); return x_4; } } +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent_goEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +lean_inc(x_1); +x_6 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_4); +if (x_6 == 0) +{ +uint8_t x_7; +lean_inc(x_1); +x_7 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_5); +if (x_7 == 0) +{ +uint8_t x_8; +lean_dec(x_1); +x_8 = 0; +return x_8; +} +else +{ +uint8_t x_9; +x_9 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_3, x_4); +return x_9; +} +} +else +{ +uint8_t x_10; +lean_inc(x_1); +x_10 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_3, x_5); +if (x_10 == 0) +{ +uint8_t x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_5); +if (x_11 == 0) +{ +uint8_t x_12; +lean_dec(x_1); +x_12 = 0; +return x_12; +} +else +{ +uint8_t x_13; +x_13 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_3, x_4); +return x_13; +} +} +else +{ +uint8_t x_14; +lean_dec(x_1); +x_14 = 1; +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent_goEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_Meta_Grind_isCongruent_goEq(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_7 = lean_box(x_6); +return x_7; +} +} LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent_go(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -4907,7 +5340,7 @@ x_4 = l_Lean_Expr_isApp(x_2); if (x_4 == 0) { uint8_t x_5; -x_5 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); +x_5 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); return x_5; @@ -4919,7 +5352,7 @@ x_6 = l_Lean_Expr_isApp(x_3); if (x_6 == 0) { uint8_t x_7; -x_7 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); +x_7 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); return x_7; @@ -4930,7 +5363,7 @@ lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = l_Lean_Expr_appArg_x21(x_2); x_9 = l_Lean_Expr_appArg_x21(x_3); lean_inc(x_1); -x_10 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_8, x_9); +x_10 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_8, x_9); lean_dec(x_9); lean_dec(x_8); if (x_10 == 0) @@ -4969,107 +5402,300 @@ return x_5; LEAN_EXPORT uint8_t l_Lean_Meta_Grind_isCongruent(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = l_Lean_Meta_Grind_congrHash___closed__2; -x_5 = lean_unsigned_to_nat(2u); -x_6 = l_Lean_Expr_isAppOfArity(x_2, x_4, x_5); -if (x_6 == 0) +lean_object* x_4; uint8_t x_5; +lean_inc(x_2); +x_4 = l_Lean_Expr_cleanupAnnotations(x_2); +x_5 = l_Lean_Expr_isApp(x_4); +if (x_5 == 0) { -uint8_t x_7; -x_7 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); -return x_7; +uint8_t x_6; +lean_dec(x_4); +x_6 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_6; } else { -uint8_t x_8; -x_8 = l_Lean_Expr_isAppOfArity(x_3, x_4, x_5); -if (x_8 == 0) +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_Expr_appArg(x_4, lean_box(0)); +x_8 = l_Lean_Expr_appFnCleanup(x_4, lean_box(0)); +x_9 = l_Lean_Expr_isApp(x_8); +if (x_9 == 0) { -uint8_t x_9; -x_9 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); -return x_9; +uint8_t x_10; +lean_dec(x_8); +lean_dec(x_7); +x_10 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_10; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_10); -x_12 = lean_nat_sub(x_11, x_10); -lean_dec(x_11); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_12, x_13); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l_Lean_Expr_appArg(x_8, lean_box(0)); +x_12 = l_Lean_Expr_appFnCleanup(x_8, lean_box(0)); +x_13 = l_Lean_Meta_Grind_congrHash___closed__2; +x_14 = l_Lean_Expr_isConstOf(x_12, x_13); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = l_Lean_Expr_isApp(x_12); +if (x_15 == 0) +{ +uint8_t x_16; lean_dec(x_12); -x_15 = l_Lean_Expr_getRevArg_x21(x_2, x_14); -lean_dec(x_2); -x_16 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_3, x_10); -x_17 = lean_nat_sub(x_16, x_10); -lean_dec(x_16); -x_18 = lean_nat_sub(x_17, x_13); -lean_dec(x_17); -x_19 = l_Lean_Expr_getRevArg_x21(x_3, x_18); -lean_dec(x_3); -x_20 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hasSameRoot(x_1, x_15, x_19); -lean_dec(x_19); -lean_dec(x_15); -return x_20; -} -} -} +lean_dec(x_11); +lean_dec(x_7); +x_16 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_16; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_Meta_Grind_isCongruent(x_1, x_2, x_3); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrKey(lean_object* x_1, lean_object* x_2) { -_start: +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = l_Lean_Expr_appArg(x_12, lean_box(0)); +x_18 = l_Lean_Expr_appFnCleanup(x_12, lean_box(0)); +x_19 = l_Lean_Meta_Grind_congrHash___closed__4; +x_20 = l_Lean_Expr_isConstOf(x_18, x_19); +lean_dec(x_18); +if (x_20 == 0) { -uint64_t x_3; -x_3 = l_Lean_Meta_Grind_congrHash(x_1, x_2); -return x_3; -} +uint8_t x_21; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +x_21 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_21; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrKey___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint64_t x_3; lean_object* x_4; -x_3 = l_Lean_Meta_Grind_instHashableCongrKey(x_1, x_2); +lean_object* x_22; uint8_t x_23; +lean_inc(x_3); +x_22 = l_Lean_Expr_cleanupAnnotations(x_3); +x_23 = l_Lean_Expr_isApp(x_22); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_3); lean_dec(x_2); -x_4 = lean_box_uint64(x_3); -return x_4; -} +lean_dec(x_1); +x_24 = 0; +return x_24; } -LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqCongrKey(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; -x_4 = l_Lean_Meta_Grind_isCongruent(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrKey___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = l_Lean_Expr_appArg(x_22, lean_box(0)); +x_26 = l_Lean_Expr_appFnCleanup(x_22, lean_box(0)); +x_27 = l_Lean_Expr_isApp(x_26); +if (x_27 == 0) { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_Meta_Grind_instBEqCongrKey(x_1, x_2, x_3); -x_5 = lean_box(x_4); -return x_5; -} +uint8_t x_28; +lean_dec(x_26); +lean_dec(x_25); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_28 = 0; +return x_28; } -LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(lean_object* x_1) { -_start: +else { -lean_object* x_2; size_t x_3; size_t x_4; size_t x_5; uint64_t x_6; -x_2 = lean_ctor_get(x_1, 0); -x_3 = lean_ptr_addr(x_2); -x_4 = 3; -x_5 = lean_usize_shift_right(x_3, x_4); -x_6 = lean_usize_to_uint64(x_5); -return x_6; +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = l_Lean_Expr_appArg(x_26, lean_box(0)); +x_30 = l_Lean_Expr_appFnCleanup(x_26, lean_box(0)); +x_31 = l_Lean_Expr_isApp(x_30); +if (x_31 == 0) +{ +uint8_t x_32; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = 0; +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = l_Lean_Expr_appArg(x_30, lean_box(0)); +x_34 = l_Lean_Expr_appFnCleanup(x_30, lean_box(0)); +x_35 = l_Lean_Expr_isConstOf(x_34, x_19); +lean_dec(x_34); +if (x_35 == 0) +{ +uint8_t x_36; +lean_dec(x_33); +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = 0; +return x_36; +} +else +{ +uint8_t x_37; +x_37 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_33); +lean_dec(x_33); +lean_dec(x_17); +if (x_37 == 0) +{ +uint8_t x_38; +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_11); +lean_dec(x_7); +x_38 = l_Lean_Meta_Grind_isCongruent_go(x_1, x_2, x_3); +return x_38; +} +else +{ +uint8_t x_39; +lean_dec(x_3); +lean_dec(x_2); +x_39 = l_Lean_Meta_Grind_isCongruent_goEq(x_1, x_11, x_7, x_29, x_25); +lean_dec(x_25); +lean_dec(x_29); +lean_dec(x_7); +lean_dec(x_11); +return x_39; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_40; uint8_t x_41; +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_2); +x_40 = l_Lean_Expr_cleanupAnnotations(x_3); +x_41 = l_Lean_Expr_isApp(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_dec(x_40); +lean_dec(x_11); +lean_dec(x_1); +x_42 = 0; +return x_42; +} +else +{ +lean_object* x_43; uint8_t x_44; +x_43 = l_Lean_Expr_appFnCleanup(x_40, lean_box(0)); +x_44 = l_Lean_Expr_isApp(x_43); +if (x_44 == 0) +{ +uint8_t x_45; +lean_dec(x_43); +lean_dec(x_11); +lean_dec(x_1); +x_45 = 0; +return x_45; +} +else +{ +lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_46 = l_Lean_Expr_appArg(x_43, lean_box(0)); +x_47 = l_Lean_Expr_appFnCleanup(x_43, lean_box(0)); +x_48 = l_Lean_Expr_isConstOf(x_47, x_13); +lean_dec(x_47); +if (x_48 == 0) +{ +uint8_t x_49; +lean_dec(x_46); +lean_dec(x_11); +lean_dec(x_1); +x_49 = 0; +return x_49; +} +else +{ +uint8_t x_50; +x_50 = l_Lean_Meta_Grind_hasSameRoot(x_1, x_11, x_46); +lean_dec(x_46); +lean_dec(x_11); +return x_50; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongruent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_Meta_Grind_isCongruent(x_1, x_2, x_3); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashableCongrKey(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; +x_3 = l_Lean_Meta_Grind_congrHash(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashableCongrKey___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_Grind_instHashableCongrKey(x_1, x_2); +x_4 = lean_box_uint64(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Grind_instBEqCongrKey(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_Meta_Grind_isCongruent(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instBEqCongrKey___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_Meta_Grind_instBEqCongrKey(x_1, x_2, x_3); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT uint64_t l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; size_t x_3; size_t x_4; size_t x_5; uint64_t x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ptr_addr(x_2); +x_4 = 3; +x_5 = lean_usize_shift_right(x_3, x_4); +x_6 = lean_usize_to_uint64(x_5); +return x_6; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1___boxed(lean_object* x_1) { @@ -5585,6 +6211,14 @@ x_1 = l_Lean_Meta_Grind_canon___closed__2; return x_1; } } +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_canon___closed__2; +return x_1; +} +} static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__1() { _start: { @@ -5633,6 +6267,19 @@ return x_5; static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_canon___closed__2; +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedEMatchTheorems___spec__1; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__6() { +_start: +{ lean_object* x_1; x_1 = l_Std_Queue_empty(lean_box(0)); return x_1; @@ -5641,33 +6288,42 @@ return x_1; static lean_object* _init_l_Lean_Meta_Grind_instInhabitedGoal() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Grind_canon___closed__2; -x_3 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1; -x_4 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; -x_5 = 0; -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Lean_Meta_Grind_instInhabitedGoal___closed__4; -x_8 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +x_2 = lean_box(0); +x_3 = l_Lean_Meta_Grind_canon___closed__2; +x_4 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1; +x_5 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; +x_6 = 0; +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Lean_Meta_Grind_instInhabitedGoal___closed__4; x_9 = l_Lean_Meta_Grind_instInhabitedGoal___closed__5; -x_10 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_2); -lean_ctor_set(x_10, 2, x_2); -lean_ctor_set(x_10, 3, x_3); -lean_ctor_set(x_10, 4, x_2); -lean_ctor_set(x_10, 5, x_4); -lean_ctor_set(x_10, 6, x_6); -lean_ctor_set(x_10, 7, x_6); -lean_ctor_set(x_10, 8, x_7); -lean_ctor_set(x_10, 9, x_7); -lean_ctor_set(x_10, 10, x_2); -lean_ctor_set(x_10, 11, x_6); -lean_ctor_set(x_10, 12, x_8); -lean_ctor_set(x_10, 13, x_9); -lean_ctor_set_uint8(x_10, sizeof(void*)*14, x_5); -return x_10; +x_10 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2; +x_11 = l_Lean_Meta_Grind_instInhabitedGoal___closed__6; +x_12 = l_Lean_PersistentHashMap_empty___at_Lean_KeyedDeclsAttribute_instInhabitedExtensionState___spec__1; +x_13 = l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3; +x_14 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_3); +lean_ctor_set(x_14, 2, x_3); +lean_ctor_set(x_14, 3, x_4); +lean_ctor_set(x_14, 4, x_3); +lean_ctor_set(x_14, 5, x_5); +lean_ctor_set(x_14, 6, x_7); +lean_ctor_set(x_14, 7, x_7); +lean_ctor_set(x_14, 8, x_8); +lean_ctor_set(x_14, 9, x_8); +lean_ctor_set(x_14, 10, x_9); +lean_ctor_set(x_14, 11, x_7); +lean_ctor_set(x_14, 12, x_7); +lean_ctor_set(x_14, 13, x_10); +lean_ctor_set(x_14, 14, x_11); +lean_ctor_set(x_14, 15, x_12); +lean_ctor_set(x_14, 16, x_1); +lean_ctor_set(x_14, 17, x_7); +lean_ctor_set(x_14, 18, x_13); +lean_ctor_set_uint8(x_14, sizeof(void*)*19, x_6); +return x_14; } } LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Goal_admit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -6113,3786 +6769,4035 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_8, 2); +x_12 = l_Lean_isTracingEnabledForCore(x_1, x_11, x_10); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -return x_6; +return x_12; } else { -lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; -x_8 = lean_array_uget(x_3, x_5); -x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); -lean_dec(x_8); -x_10 = lean_uint64_mix_hash(x_6, x_9); -x_11 = 1; -x_12 = lean_usize_add(x_5, x_11); -x_5 = x_12; -x_6 = x_10; -goto _start; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } } -LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +static double _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1() { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -return x_6; +lean_object* x_1; uint8_t x_2; double x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = l_Float_ofScientific(x_1, x_2, x_1); +return x_3; } -else -{ -lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; -x_8 = lean_array_uget(x_3, x_5); -x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); -lean_dec(x_8); -x_10 = lean_uint64_mix_hash(x_6, x_9); -x_11 = 1; -x_12 = lean_usize_add(x_5, x_11); -x_5 = x_12; -x_6 = x_10; -goto _start; } +static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3() { _start: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_5); -return x_6; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_9; lean_object* x_10; uint64_t x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; uint64_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_9); -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -x_13 = lean_box(0); -x_14 = lean_array_size(x_12); -x_15 = 0; -x_16 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_12, x_13, x_12, x_14, x_15, x_11); -lean_dec(x_12); -x_17 = lean_uint64_to_usize(x_16); -x_18 = 1; -x_19 = lean_usize_sub(x_1, x_18); -x_20 = 5; -x_21 = lean_usize_mul(x_20, x_19); -x_22 = lean_usize_shift_right(x_17, x_21); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_add(x_5, x_23); -lean_dec(x_5); -x_25 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_6, x_22, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_24; -x_6 = x_25; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) -{ -lean_dec(x_3); -return x_7; -} -else -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_ctor_get(x_7, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = lean_ctor_get(x_9, 5); +x_13 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_7, x_8, x_9, x_10, x_11); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 2); +x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 0) -{ -lean_dec(x_15); -lean_dec(x_14); lean_dec(x_13); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_3); -return x_7; -} -else -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_11); -if (x_17 == 0) +x_16 = lean_st_ref_take(x_10, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 3); +lean_inc(x_18); +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_18 = lean_ctor_get(x_11, 2); -lean_dec(x_18); -x_19 = lean_ctor_get(x_11, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_array_fget(x_13, x_14); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_14, x_22); -lean_dec(x_14); -lean_ctor_set(x_11, 1, x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_16, 1); +x_21 = lean_ctor_get(x_16, 0); lean_dec(x_21); -lean_dec(x_9); +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_17, 3); +lean_dec(x_23); +x_24 = !lean_is_exclusive(x_18); if (x_24 == 0) { -lean_object* x_25; -lean_dec(x_3); -x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 0, x_25); -return x_7; +lean_object* x_25; double x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_25 = lean_ctor_get(x_18, 0); +x_26 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; +x_27 = 0; +x_28 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_29 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_29, 0, x_1); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set_float(x_29, sizeof(void*)*2, x_26); +lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_26); +lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_27); +x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_31 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_14); +lean_ctor_set(x_31, 2, x_30); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_31); +lean_ctor_set(x_16, 0, x_12); +x_32 = l_Lean_PersistentArray_push___rarg(x_25, x_16); +lean_ctor_set(x_18, 0, x_32); +x_33 = lean_st_ref_set(x_10, x_17, x_20); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +x_36 = lean_box(0); +lean_ctor_set(x_33, 0, x_36); +return x_33; } else { -size_t x_26; size_t x_27; -lean_inc(x_3); -lean_ctor_set(x_7, 0, x_3); -x_26 = 1; -x_27 = lean_usize_add(x_6, x_26); -x_6 = x_27; -goto _start; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +return x_39; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_dec(x_11); -x_29 = lean_array_fget(x_13, x_14); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_14, x_30); -lean_dec(x_14); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_15); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); -lean_dec(x_29); -lean_dec(x_9); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_3); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_34); -return x_7; +uint64_t x_40; lean_object* x_41; double x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_40 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_41 = lean_ctor_get(x_18, 0); +lean_inc(x_41); +lean_dec(x_18); +x_42 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; +x_43 = 0; +x_44 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_45 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_45, 0, x_1); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set_float(x_45, sizeof(void*)*2, x_42); +lean_ctor_set_float(x_45, sizeof(void*)*2 + 8, x_42); +lean_ctor_set_uint8(x_45, sizeof(void*)*2 + 16, x_43); +x_46 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_47 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_14); +lean_ctor_set(x_47, 2, x_46); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_47); +lean_ctor_set(x_16, 0, x_12); +x_48 = l_Lean_PersistentArray_push___rarg(x_41, x_16); +x_49 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set_uint64(x_49, sizeof(void*)*1, x_40); +lean_ctor_set(x_17, 3, x_49); +x_50 = lean_st_ref_set(x_10, x_17, x_20); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_52 = x_50; +} else { + lean_dec_ref(x_50); + x_52 = lean_box(0); +} +x_53 = lean_box(0); +if (lean_is_scalar(x_52)) { + x_54 = lean_alloc_ctor(0, 2, 0); +} else { + x_54 = x_52; +} +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_51); +return x_54; +} } else { -size_t x_35; size_t x_36; -lean_inc(x_3); -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_3); -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -goto _start; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint64_t x_62; lean_object* x_63; lean_object* x_64; double x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_55 = lean_ctor_get(x_17, 0); +x_56 = lean_ctor_get(x_17, 1); +x_57 = lean_ctor_get(x_17, 2); +x_58 = lean_ctor_get(x_17, 4); +x_59 = lean_ctor_get(x_17, 5); +x_60 = lean_ctor_get(x_17, 6); +x_61 = lean_ctor_get(x_17, 7); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_17); +x_62 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_63 = lean_ctor_get(x_18, 0); +lean_inc(x_63); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + x_64 = x_18; +} else { + lean_dec_ref(x_18); + x_64 = lean_box(0); +} +x_65 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; +x_66 = 0; +x_67 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_68 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_68, 0, x_1); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set_float(x_68, sizeof(void*)*2, x_65); +lean_ctor_set_float(x_68, sizeof(void*)*2 + 8, x_65); +lean_ctor_set_uint8(x_68, sizeof(void*)*2 + 16, x_66); +x_69 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_70 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_14); +lean_ctor_set(x_70, 2, x_69); +lean_inc(x_12); +lean_ctor_set(x_16, 1, x_70); +lean_ctor_set(x_16, 0, x_12); +x_71 = l_Lean_PersistentArray_push___rarg(x_63, x_16); +if (lean_is_scalar(x_64)) { + x_72 = lean_alloc_ctor(0, 1, 8); +} else { + x_72 = x_64; +} +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set_uint64(x_72, sizeof(void*)*1, x_62); +x_73 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_73, 0, x_55); +lean_ctor_set(x_73, 1, x_56); +lean_ctor_set(x_73, 2, x_57); +lean_ctor_set(x_73, 3, x_72); +lean_ctor_set(x_73, 4, x_58); +lean_ctor_set(x_73, 5, x_59); +lean_ctor_set(x_73, 6, x_60); +lean_ctor_set(x_73, 7, x_61); +x_74 = lean_st_ref_set(x_10, x_73, x_20); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_76 = x_74; +} else { + lean_dec_ref(x_74); + x_76 = lean_box(0); } +x_77 = lean_box(0); +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(0, 2, 0); +} else { + x_78 = x_76; } +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_75); +return x_78; } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_7, 1); -lean_inc(x_38); -lean_dec(x_7); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 2); -lean_inc(x_41); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 0) -{ -lean_object* x_43; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_9); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_3); -lean_ctor_set(x_43, 1, x_38); -return x_43; +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint64_t x_88; lean_object* x_89; lean_object* x_90; double x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_79 = lean_ctor_get(x_16, 1); +lean_inc(x_79); +lean_dec(x_16); +x_80 = lean_ctor_get(x_17, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_17, 1); +lean_inc(x_81); +x_82 = lean_ctor_get(x_17, 2); +lean_inc(x_82); +x_83 = lean_ctor_get(x_17, 4); +lean_inc(x_83); +x_84 = lean_ctor_get(x_17, 5); +lean_inc(x_84); +x_85 = lean_ctor_get(x_17, 6); +lean_inc(x_85); +x_86 = lean_ctor_get(x_17, 7); +lean_inc(x_86); +if (lean_is_exclusive(x_17)) { + lean_ctor_release(x_17, 0); + lean_ctor_release(x_17, 1); + lean_ctor_release(x_17, 2); + lean_ctor_release(x_17, 3); + lean_ctor_release(x_17, 4); + lean_ctor_release(x_17, 5); + lean_ctor_release(x_17, 6); + lean_ctor_release(x_17, 7); + x_87 = x_17; +} else { + lean_dec_ref(x_17); + x_87 = lean_box(0); +} +x_88 = lean_ctor_get_uint64(x_18, sizeof(void*)*1); +x_89 = lean_ctor_get(x_18, 0); +lean_inc(x_89); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + x_90 = x_18; +} else { + lean_dec_ref(x_18); + x_90 = lean_box(0); +} +x_91 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1; +x_92 = 0; +x_93 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_94 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_94, 0, x_1); +lean_ctor_set(x_94, 1, x_93); +lean_ctor_set_float(x_94, sizeof(void*)*2, x_91); +lean_ctor_set_float(x_94, sizeof(void*)*2 + 8, x_91); +lean_ctor_set_uint8(x_94, sizeof(void*)*2 + 16, x_92); +x_95 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_96 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_14); +lean_ctor_set(x_96, 2, x_95); +lean_inc(x_12); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_12); +lean_ctor_set(x_97, 1, x_96); +x_98 = l_Lean_PersistentArray_push___rarg(x_89, x_97); +if (lean_is_scalar(x_90)) { + x_99 = lean_alloc_ctor(0, 1, 8); +} else { + x_99 = x_90; } -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_44 = x_38; +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set_uint64(x_99, sizeof(void*)*1, x_88); +if (lean_is_scalar(x_87)) { + x_100 = lean_alloc_ctor(0, 8, 0); } else { - lean_dec_ref(x_38); - x_44 = lean_box(0); + x_100 = x_87; +} +lean_ctor_set(x_100, 0, x_80); +lean_ctor_set(x_100, 1, x_81); +lean_ctor_set(x_100, 2, x_82); +lean_ctor_set(x_100, 3, x_99); +lean_ctor_set(x_100, 4, x_83); +lean_ctor_set(x_100, 5, x_84); +lean_ctor_set(x_100, 6, x_85); +lean_ctor_set(x_100, 7, x_86); +x_101 = lean_st_ref_set(x_10, x_100, x_79); +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_103 = x_101; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); } -x_45 = lean_array_fget(x_39, x_40); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_40, x_46); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_48 = lean_alloc_ctor(0, 3, 0); +x_104 = lean_box(0); +if (lean_is_scalar(x_103)) { + x_105 = lean_alloc_ctor(0, 2, 0); } else { - x_48 = x_44; + x_105 = x_103; } -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_41); -x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); -lean_dec(x_45); -lean_dec(x_9); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -lean_dec(x_3); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_102); +return x_105; } -else +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_52; size_t x_53; size_t x_54; -lean_inc(x_3); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_48); -x_53 = 1; -x_54 = lean_usize_add(x_6, x_53); -x_6 = x_54; -x_7 = x_52; -goto _start; +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_6, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_13, 7); +lean_dec(x_16); +lean_ctor_set(x_13, 7, x_1); +x_17 = lean_st_ref_set(x_6, x_13, x_14); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_17, 0, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; } } +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +x_26 = lean_ctor_get(x_13, 2); +x_27 = lean_ctor_get(x_13, 3); +x_28 = lean_ctor_get(x_13, 4); +x_29 = lean_ctor_get(x_13, 5); +x_30 = lean_ctor_get(x_13, 6); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_31 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_31, 0, x_24); +lean_ctor_set(x_31, 1, x_25); +lean_ctor_set(x_31, 2, x_26); +lean_ctor_set(x_31, 3, x_27); +lean_ctor_set(x_31, 4, x_28); +lean_ctor_set(x_31, 5, x_29); +lean_ctor_set(x_31, 6, x_30); +lean_ctor_set(x_31, 7, x_1); +x_32 = lean_st_ref_set(x_6, x_31, x_14); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_34 = x_32; +} else { + lean_dec_ref(x_32); + x_34 = lean_box(0); +} +x_35 = lean_box(0); +if (lean_is_scalar(x_34)) { + x_36 = lean_alloc_ctor(0, 2, 0); +} else { + x_36 = x_34; } +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_33); +return x_36; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__1() { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__2() { +_start: { -lean_dec(x_3); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("working on goal `", 17, 17); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__3() { +_start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_updateLastTag___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__4() { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_ctor_get(x_7, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 2); -lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_updateLastTag___closed__5() { +_start: { -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_3); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_updateLastTag___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_11); -if (x_17 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = l_Lean_Meta_Grind_updateLastTag___closed__1; +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_18 = lean_ctor_get(x_11, 2); -lean_dec(x_18); -x_19 = lean_ctor_get(x_11, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_array_fget(x_13, x_14); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_14, x_22); -lean_dec(x_14); -lean_ctor_set(x_11, 1, x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); -lean_dec(x_21); -lean_dec(x_9); -if (x_24 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) { -lean_object* x_25; -lean_dec(x_3); -x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 0, x_25); -return x_7; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_11, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); +return x_11; } else { -size_t x_26; size_t x_27; -lean_inc(x_3); -lean_ctor_set(x_7, 0, x_3); -x_26 = 1; -x_27 = lean_usize_add(x_6, x_26); -x_6 = x_27; -goto _start; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); lean_dec(x_11); -x_29 = lean_array_fget(x_13, x_14); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_14, x_30); -lean_dec(x_14); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_15); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); -lean_dec(x_29); -lean_dec(x_9); -if (x_33 == 0) +x_21 = lean_st_ref_get(x_1, x_20); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_34; -lean_dec(x_3); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_34); -return x_7; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_ctor_get(x_21, 1); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_MVarId_getTag(x_25, x_5, x_6, x_7, x_8, x_24); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_st_ref_get(x_4, x_28); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +x_33 = lean_ctor_get(x_31, 7); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_name_eq(x_27, x_33); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; +lean_free_object(x_29); +x_35 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_unbox(x_36); +lean_dec(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_free_object(x_21); +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_dec(x_35); +x_39 = lean_box(0); +x_40 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_39, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_38); +return x_40; } else { -size_t x_35; size_t x_36; -lean_inc(x_3); -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_3); -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -goto _start; +uint8_t x_41; +x_41 = !lean_is_exclusive(x_35); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_42 = lean_ctor_get(x_35, 1); +x_43 = lean_ctor_get(x_35, 0); +lean_dec(x_43); +lean_inc(x_27); +x_44 = l_Lean_MessageData_ofName(x_27); +x_45 = l_Lean_Meta_Grind_updateLastTag___closed__3; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_44); +lean_ctor_set(x_35, 0, x_45); +x_46 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_21, 7); +lean_ctor_set(x_21, 1, x_46); +lean_ctor_set(x_21, 0, x_35); +x_47 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_10, x_21, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_42); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_48, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_49); +lean_dec(x_48); +return x_50; } +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_51 = lean_ctor_get(x_35, 1); +lean_inc(x_51); +lean_dec(x_35); +lean_inc(x_27); +x_52 = l_Lean_MessageData_ofName(x_27); +x_53 = l_Lean_Meta_Grind_updateLastTag___closed__3; +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); +x_55 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_21, 7); +lean_ctor_set(x_21, 1, x_55); +lean_ctor_set(x_21, 0, x_54); +x_56 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_10, x_21, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_51); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_57, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_58); +lean_dec(x_57); +return x_59; } } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_7, 1); -lean_inc(x_38); -lean_dec(x_7); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 2); -lean_inc(x_41); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 0) +lean_object* x_60; +lean_dec(x_27); +lean_free_object(x_21); +x_60 = lean_box(0); +lean_ctor_set(x_29, 0, x_60); +return x_29; +} +} +else { -lean_object* x_43; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_9); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_3); -lean_ctor_set(x_43, 1, x_38); -return x_43; +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_61 = lean_ctor_get(x_29, 0); +x_62 = lean_ctor_get(x_29, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_29); +x_63 = lean_ctor_get(x_61, 7); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_name_eq(x_27, x_63); +lean_dec(x_63); +if (x_64 == 0) +{ +lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_65 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_62); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_unbox(x_66); +lean_dec(x_66); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_free_object(x_21); +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +lean_dec(x_65); +x_69 = lean_box(0); +x_70 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_69, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_68); +return x_70; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_44 = x_38; +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_71 = lean_ctor_get(x_65, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_72 = x_65; } else { - lean_dec_ref(x_38); - x_44 = lean_box(0); + lean_dec_ref(x_65); + x_72 = lean_box(0); } -x_45 = lean_array_fget(x_39, x_40); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_40, x_46); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_48 = lean_alloc_ctor(0, 3, 0); +lean_inc(x_27); +x_73 = l_Lean_MessageData_ofName(x_27); +x_74 = l_Lean_Meta_Grind_updateLastTag___closed__3; +if (lean_is_scalar(x_72)) { + x_75 = lean_alloc_ctor(7, 2, 0); } else { - x_48 = x_44; + x_75 = x_72; + lean_ctor_set_tag(x_75, 7); +} +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_73); +x_76 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_21, 7); +lean_ctor_set(x_21, 1, x_76); +lean_ctor_set(x_21, 0, x_75); +x_77 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_10, x_21, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_71); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +lean_dec(x_77); +x_80 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_27, x_78, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_79); +lean_dec(x_78); +return x_80; } -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_41); -x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); -lean_dec(x_45); -lean_dec(x_9); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -lean_dec(x_3); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; } else { -lean_object* x_52; size_t x_53; size_t x_54; -lean_inc(x_3); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_48); -x_53 = 1; -x_54 = lean_usize_add(x_6, x_53); -x_6 = x_54; -x_7 = x_52; -goto _start; -} -} -} +lean_object* x_81; lean_object* x_82; +lean_dec(x_27); +lean_free_object(x_21); +x_81 = lean_box(0); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_62); +return x_82; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_box(0); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_7); -x_12 = lean_array_size(x_8); -x_13 = 0; -x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_8, x_9, x_10, x_8, x_12, x_13, x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +uint8_t x_83; +lean_free_object(x_21); +x_83 = !lean_is_exclusive(x_26); +if (x_83 == 0) { -uint8_t x_16; -x_16 = 1; -return x_16; +return x_26; } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -return x_18; +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_26, 0); +x_85 = lean_ctor_get(x_26, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_26); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_dec_eq(x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); -if (x_8 == 0) +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_87 = lean_ctor_get(x_21, 0); +x_88 = lean_ctor_get(x_21, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_21); +x_89 = lean_ctor_get(x_87, 0); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_MVarId_getTag(x_89, x_5, x_6, x_7, x_8, x_88); +if (lean_obj_tag(x_90) == 0) { -uint8_t x_9; -lean_dec(x_1); -x_9 = 0; -return x_9; +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = lean_st_ref_get(x_4, x_92); +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_93, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_96 = x_93; +} else { + lean_dec_ref(x_93); + x_96 = lean_box(0); +} +x_97 = lean_ctor_get(x_94, 7); +lean_inc(x_97); +lean_dec(x_94); +x_98 = lean_name_eq(x_91, x_97); +lean_dec(x_97); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; uint8_t x_101; +lean_dec(x_96); +x_99 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_95); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_unbox(x_100); +lean_dec(x_100); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_99, 1); +lean_inc(x_102); +lean_dec(x_99); +x_103 = lean_box(0); +x_104 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_91, x_103, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_102); +return x_104; } else { -lean_object* x_10; uint8_t x_11; -x_10 = lean_box(0); -x_11 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_10); -return x_11; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_105 = lean_ctor_get(x_99, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_106 = x_99; +} else { + lean_dec_ref(x_99); + x_106 = lean_box(0); } +lean_inc(x_91); +x_107 = l_Lean_MessageData_ofName(x_91); +x_108 = l_Lean_Meta_Grind_updateLastTag___closed__3; +if (lean_is_scalar(x_106)) { + x_109 = lean_alloc_ctor(7, 2, 0); +} else { + x_109 = x_106; + lean_ctor_set_tag(x_109, 7); } +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_107); +x_110 = l_Lean_Meta_Grind_updateLastTag___closed__5; +x_111 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_10, x_111, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_105); +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +lean_dec(x_112); +x_115 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_91, x_113, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_114); +lean_dec(x_113); +return x_115; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_1, 1); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 0); -lean_dec(x_11); -x_12 = lean_array_push(x_5, x_3); -x_13 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_12); -return x_1; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_1); -x_14 = lean_array_push(x_5, x_3); -x_15 = lean_array_push(x_6, x_4); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_object* x_116; lean_object* x_117; +lean_dec(x_91); +x_116 = lean_box(0); +if (lean_is_scalar(x_96)) { + x_117 = lean_alloc_ctor(0, 2, 0); +} else { + x_117 = x_96; +} +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_95); +return x_117; } } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_ctor_get(x_3, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -x_20 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_18, x_19); -lean_dec(x_19); -lean_dec(x_18); -if (x_20 == 0) +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_118 = lean_ctor_get(x_90, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_90, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_120 = x_90; +} else { + lean_dec_ref(x_90); + x_120 = lean_box(0); +} +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); +} else { + x_121 = x_120; +} +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_21; lean_object* x_22; -lean_dec(x_17); +lean_object* x_11; +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_add(x_2, x_21); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_2 = x_22; -goto _start; +return x_11; } -else -{ -lean_object* x_24; uint8_t x_25; -x_24 = lean_box(0); -x_25 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_17, x_3, x_24); -if (x_25 == 0) +} +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_26; lean_object* x_27; +lean_object* x_12; +x_12 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_26 = lean_unsigned_to_nat(1u); -x_27 = lean_nat_add(x_2, x_26); -lean_dec(x_2); -x_2 = x_27; -goto _start; +lean_dec(x_4); +lean_dec(x_3); +return x_12; } -else -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_1); -if (x_29 == 0) +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_1, 1); -lean_dec(x_30); -x_31 = lean_ctor_get(x_1, 0); -lean_dec(x_31); -x_32 = lean_array_fset(x_5, x_2, x_3); -x_33 = lean_array_fset(x_6, x_2, x_4); +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_updateLastTag___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -lean_ctor_set(x_1, 1, x_33); -lean_ctor_set(x_1, 0, x_32); -return x_1; +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_updateLastTag___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_1); -x_34 = lean_array_fset(x_5, x_2, x_3); -x_35 = lean_array_fset(x_6, x_2, x_4); +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_updateLastTag(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_dec(x_1); +return x_10; } } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doElemTrace_goal[_]__", 21, 21); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__7; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__8; +x_4 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_box(0); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_7); -x_12 = lean_array_size(x_8); -x_13 = 0; -x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_8, x_9, x_10, x_8, x_12, x_13, x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = 1; -return x_16; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; } -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -return x_18; } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_dec_eq(x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_1); -x_9 = 0; -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("trace_goal[", 11, 11); +return x_1; } -else -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_box(0); -x_11 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_10); -return x_11; } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7() { _start: { -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = 1; -x_9 = 5; -x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_14 = lean_nat_dec_lt(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ident", 5, 5); return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8() { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9() { +_start: { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10() { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_ctor_get(x_4, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 0); -lean_inc(x_22); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_22); -lean_dec(x_22); -lean_dec(x_21); -if (x_23 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6; +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11() { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_free_object(x_15); -x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = lean_array_fset(x_17, x_12, x_25); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_26); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12() { +_start: { -uint8_t x_27; -lean_inc(x_19); -x_27 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_19, x_4, x_16); -if (x_27 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13() { +_start: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_free_object(x_15); -x_28 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = lean_array_fset(x_17, x_12, x_29); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_30); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10; +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14() { +_start: { -lean_object* x_31; -lean_dec(x_20); -lean_dec(x_19); -lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_4); -x_31 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("orelse", 6, 6); return x_1; } } -} -else +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_32 = lean_ctor_get(x_15, 0); -x_33 = lean_ctor_get(x_15, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_15); -x_34 = lean_ctor_get(x_4, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_32, 0); -lean_inc(x_35); -x_36 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_34, x_35); -lean_dec(x_35); -lean_dec(x_34); -if (x_36 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16() { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_array_fset(x_17, x_12, x_38); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("interpolatedStr", 15, 15); return x_1; } -else -{ -uint8_t x_40; -lean_inc(x_32); -x_40 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_32, x_4, x_16); -if (x_40 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17() { +_start: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_43 = lean_array_fset(x_17, x_12, x_42); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_43); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18() { +_start: { -lean_object* x_44; lean_object* x_45; -lean_dec(x_33); -lean_dec(x_32); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_4); -lean_ctor_set(x_44, 1, x_5); -x_45 = lean_array_fset(x_17, x_12, x_44); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term", 4, 4); return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -case 1: +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20() { +_start: { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_15); -if (x_46 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21() { +_start: { -lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; -x_47 = lean_ctor_get(x_15, 0); -x_48 = lean_usize_shift_right(x_2, x_9); -x_49 = lean_usize_add(x_3, x_8); -x_50 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_47, x_48, x_49, x_4, x_5); -lean_ctor_set(x_15, 0, x_50); -x_51 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_51); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22() { +_start: { -lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_15, 0); -lean_inc(x_52); -lean_dec(x_15); -x_53 = lean_usize_shift_right(x_2, x_9); -x_54 = lean_usize_add(x_3, x_8); -x_55 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_52, x_53, x_54, x_4, x_5); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_55); -x_57 = lean_array_fset(x_17, x_12, x_56); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_57); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21; +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -default: +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23() { +_start: { -lean_object* x_58; lean_object* x_59; -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_4); -lean_ctor_set(x_58, 1, x_5); -x_59 = lean_array_fset(x_17, x_12, x_58); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_59); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4; +x_2 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13; +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -else +static lean_object* _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____() { +_start: { -lean_object* x_60; size_t x_61; size_t x_62; size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_60 = lean_ctor_get(x_1, 0); -lean_inc(x_60); -lean_dec(x_1); -x_61 = 1; -x_62 = 5; -x_63 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_64 = lean_usize_land(x_2, x_63); -x_65 = lean_usize_to_nat(x_64); -x_66 = lean_array_get_size(x_60); -x_67 = lean_nat_dec_lt(x_65, x_66); -lean_dec(x_66); -if (x_67 == 0) +lean_object* x_1; +x_1 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1() { +_start: { -lean_object* x_68; -lean_dec(x_65); -lean_dec(x_5); -lean_dec(x_4); -x_68 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_68, 0, x_60); -return x_68; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2() { +_start: { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_array_fget(x_60, x_65); -x_70 = lean_box(0); -x_71 = lean_array_fset(x_60, x_65, x_70); -switch (lean_obj_tag(x_69)) { -case 0: +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3() { +_start: { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_72 = lean_ctor_get(x_69, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_69, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - lean_ctor_release(x_69, 1); - x_74 = x_69; -} else { - lean_dec_ref(x_69); - x_74 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doNested", 8, 8); +return x_1; } -x_75 = lean_ctor_get(x_4, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_72, 0); -lean_inc(x_76); -x_77 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_75, x_76); -lean_dec(x_76); -lean_dec(x_75); -if (x_77 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4() { +_start: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -lean_dec(x_74); -x_78 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); -x_79 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_79, 0, x_78); -x_80 = lean_array_fset(x_71, x_65, x_79); -lean_dec(x_65); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5() { +_start: { -uint8_t x_82; -lean_inc(x_72); -x_82 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_72, x_4, x_70); -if (x_82 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("do", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6() { +_start: { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_74); -x_83 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); -x_84 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_84, 0, x_83); -x_85 = lean_array_fset(x_71, x_65, x_84); -lean_dec(x_65); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -return x_86; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doSeqIndent", 11, 11); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7() { +_start: { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_dec(x_73); -lean_dec(x_72); -if (lean_is_scalar(x_74)) { - x_87 = lean_alloc_ctor(0, 2, 0); -} else { - x_87 = x_74; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -lean_ctor_set(x_87, 0, x_4); -lean_ctor_set(x_87, 1, x_5); -x_88 = lean_array_fset(x_71, x_65, x_87); -lean_dec(x_65); -x_89 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_89, 0, x_88); -return x_89; } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; } } -case 1: +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9() { +_start: { -lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_90 = lean_ctor_get(x_69, 0); -lean_inc(x_90); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - x_91 = x_69; -} else { - lean_dec_ref(x_69); - x_91 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -x_92 = lean_usize_shift_right(x_2, x_62); -x_93 = lean_usize_add(x_3, x_61); -x_94 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_90, x_92, x_93, x_4, x_5); -if (lean_is_scalar(x_91)) { - x_95 = lean_alloc_ctor(1, 1, 0); -} else { - x_95 = x_91; } -lean_ctor_set(x_95, 0, x_94); -x_96 = lean_array_fset(x_71, x_65, x_95); -lean_dec(x_65); -x_97 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_97, 0, x_96); -return x_97; +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doSeqItem", 9, 9); +return x_1; } -default: +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11() { +_start: { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_4); -lean_ctor_set(x_98, 1, x_5); -x_99 = lean_array_fset(x_71, x_65, x_98); -lean_dec(x_65); -x_100 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_100, 0, x_99); -return x_100; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doLet", 5, 5); +return x_1; +} } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14() { +_start: { -uint8_t x_101; -x_101 = !lean_is_exclusive(x_1); -if (x_101 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("let", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15() { +_start: { -lean_object* x_102; lean_object* x_103; size_t x_104; uint8_t x_105; -x_102 = lean_unsigned_to_nat(0u); -x_103 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_1, x_102, x_4, x_5); -x_104 = 7; -x_105 = lean_usize_dec_le(x_104, x_3); -if (x_105 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("letDecl", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16() { +_start: { -lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_106 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_103); -x_107 = lean_unsigned_to_nat(4u); -x_108 = lean_nat_dec_lt(x_106, x_107); -lean_dec(x_106); -if (x_108 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17() { +_start: { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_109 = lean_ctor_get(x_103, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_103, 1); -lean_inc(x_110); -lean_dec(x_103); -x_111 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_112 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_109, x_110, lean_box(0), x_102, x_111); -lean_dec(x_110); -lean_dec(x_109); -return x_112; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("letIdDecl", 9, 9); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18() { +_start: { -return x_103; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19() { +_start: { -return x_103; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cls", 3, 3); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20() { +_start: { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; size_t x_118; uint8_t x_119; -x_113 = lean_ctor_get(x_1, 0); -x_114 = lean_ctor_get(x_1, 1); -lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_1); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -x_116 = lean_unsigned_to_nat(0u); -x_117 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_115, x_116, x_4, x_5); -x_118 = 7; -x_119 = lean_usize_dec_le(x_118, x_3); -if (x_119 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21() { +_start: { -lean_object* x_120; lean_object* x_121; uint8_t x_122; -x_120 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_117); -x_121 = lean_unsigned_to_nat(4u); -x_122 = lean_nat_dec_lt(x_120, x_121); -lean_dec(x_120); -if (x_122 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22() { +_start: { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_123 = lean_ctor_get(x_117, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_117, 1); -lean_inc(x_124); -lean_dec(x_117); -x_125 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_126 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_123, x_124, lean_box(0), x_116, x_125); -lean_dec(x_124); -lean_dec(x_123); -return x_126; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":=", 2, 2); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23() { +_start: { -return x_117; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doIf", 4, 4); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24() { +_start: { -return x_117; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("if", 2, 2); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doIfProp", 8, 8); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27() { _start: { -uint64_t x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint64_t x_9; size_t x_10; size_t x_11; lean_object* x_12; -x_4 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); -x_5 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -x_6 = lean_box(0); -x_7 = lean_array_size(x_5); -x_8 = 0; -x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_5, x_6, x_5, x_7, x_8, x_4); -lean_dec(x_5); -x_10 = lean_uint64_to_usize(x_9); -x_11 = 1; -x_12 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_1, x_10, x_11, x_2, x_3); -return x_12; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28() { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("paren", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29() { +_start: { -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30() { +_start: { -lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; -x_8 = lean_array_uget(x_3, x_5); -x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); -lean_dec(x_8); -x_10 = lean_uint64_mix_hash(x_6, x_9); -x_11 = 1; -x_12 = lean_usize_add(x_5, x_11); -x_5 = x_12; -x_6 = x_10; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("liftMethod", 10, 10); +return x_1; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32() { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33() { +_start: { -lean_dec(x_3); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("←", 3, 1); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34() { +_start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("app", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35() { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_ctor_get(x_7, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 2); -lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36() { +_start: { -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_3); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.isTracingEnabledFor", 24, 24); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37() { +_start: { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_11); -if (x_17 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38() { +_start: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_18 = lean_ctor_get(x_11, 2); -lean_dec(x_18); -x_19 = lean_ctor_get(x_11, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_array_fget(x_13, x_14); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_14, x_22); -lean_dec(x_14); -lean_ctor_set(x_11, 1, x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); -lean_dec(x_21); -lean_dec(x_9); -if (x_24 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isTracingEnabledFor", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39() { +_start: { -lean_object* x_25; -lean_dec(x_3); -x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 0, x_25); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40() { +_start: { -size_t x_26; size_t x_27; -lean_inc(x_3); -lean_ctor_set(x_7, 0, x_3); -x_26 = 1; -x_27 = lean_usize_add(x_6, x_26); -x_6 = x_27; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41() { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_dec(x_11); -x_29 = lean_array_fget(x_13, x_14); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_14, x_30); -lean_dec(x_14); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_15); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); -lean_dec(x_29); -lean_dec(x_9); -if (x_33 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42() { +_start: { -lean_object* x_34; -lean_dec(x_3); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_34); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43() { +_start: { -size_t x_35; size_t x_36; -lean_inc(x_3); -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_3); -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("then", 4, 4); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doExpr", 6, 6); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45() { +_start: { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_7, 1); -lean_inc(x_38); -lean_dec(x_7); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 2); -lean_inc(x_41); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46() { +_start: { -lean_object* x_43; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_9); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_3); -lean_ctor_set(x_43, 1, x_38); -return x_43; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("updateLastTag", 13, 13); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47() { +_start: { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_44 = x_38; -} else { - lean_dec_ref(x_38); - x_44 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; } -x_45 = lean_array_fget(x_39, x_40); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_40, x_46); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_48 = lean_alloc_ctor(0, 3, 0); -} else { - x_48 = x_44; } -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_41); -x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); -lean_dec(x_45); -lean_dec(x_9); -if (x_49 == 0) +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48() { +_start: { -lean_object* x_50; lean_object* x_51; -lean_dec(x_3); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49() { +_start: { -lean_object* x_52; size_t x_53; size_t x_54; -lean_inc(x_3); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_48); -x_53 = 1; -x_54 = lean_usize_add(x_6, x_53); -x_6 = x_54; -x_7 = x_52; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__7; +x_3 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__8; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52() { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.addTrace", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53() { +_start: { -lean_dec(x_3); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54() { +_start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("addTrace", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55() { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_ctor_get(x_7, 0); -lean_dec(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_11, 2); -lean_inc(x_15); -x_16 = lean_nat_dec_lt(x_14, x_15); -if (x_16 == 0) -{ -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_3); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_11); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_18 = lean_ctor_get(x_11, 2); -lean_dec(x_18); -x_19 = lean_ctor_get(x_11, 1); -lean_dec(x_19); -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_array_fget(x_13, x_14); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_14, x_22); -lean_dec(x_14); -lean_ctor_set(x_11, 1, x_23); -x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); -lean_dec(x_21); -lean_dec(x_9); -if (x_24 == 0) -{ -lean_object* x_25; -lean_dec(x_3); -x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 0, x_25); -return x_7; } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56() { +_start: { -size_t x_26; size_t x_27; -lean_inc(x_3); -lean_ctor_set(x_7, 0, x_3); -x_26 = 1; -x_27 = lean_usize_add(x_6, x_26); -x_6 = x_27; -goto _start; -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_dec(x_11); -x_29 = lean_array_fget(x_13, x_14); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_14, x_30); -lean_dec(x_14); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_15); -x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); -lean_dec(x_29); -lean_dec(x_9); -if (x_33 == 0) -{ -lean_object* x_34; -lean_dec(x_3); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_34); -return x_7; } -else +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57() { +_start: { -size_t x_35; size_t x_36; -lean_inc(x_3); -lean_ctor_set(x_7, 1, x_32); -lean_ctor_set(x_7, 0, x_3); -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("quotedName", 10, 10); +return x_1; } } -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_7, 1); -lean_inc(x_38); -lean_dec(x_7); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 2); -lean_inc(x_41); -x_42 = lean_nat_dec_lt(x_40, x_41); -if (x_42 == 0) +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59() { +_start: { -lean_object* x_43; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_9); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_3); -lean_ctor_set(x_43, 1, x_38); -return x_43; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60() { +_start: { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_44 = x_38; -} else { - lean_dec_ref(x_38); - x_44 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; } -x_45 = lean_array_fget(x_39, x_40); -x_46 = lean_unsigned_to_nat(1u); -x_47 = lean_nat_add(x_40, x_46); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_48 = lean_alloc_ctor(0, 3, 0); -} else { - x_48 = x_44; } -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_41); -x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); -lean_dec(x_45); -lean_dec(x_9); -if (x_49 == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_50; lean_object* x_51; +lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_5 = lean_ctor_get(x_3, 5); +lean_inc(x_5); +x_6 = 0; +x_7 = l_Lean_SourceInfo_fromRef(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_3, 2); +lean_inc(x_8); +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); lean_dec(x_3); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; +x_10 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5; +lean_inc(x_7); +x_11 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +x_12 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14; +lean_inc(x_7); +x_13 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_13, 0, x_7); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9; +x_15 = l_Lean_Meta_Grind_instInhabitedGoal___closed__1; +lean_inc(x_7); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 2, x_15); +x_17 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21; +lean_inc(x_8); +lean_inc(x_9); +x_18 = l_Lean_addMacroScope(x_9, x_17, x_8); +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20; +lean_inc(x_7); +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_7); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_19); +x_22 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22; +lean_inc(x_7); +x_23 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_23, 0, x_7); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_Syntax_getId(x_1); +x_25 = lean_erase_macro_scopes(x_24); +lean_inc(x_25); +x_26 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_25); +x_27 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25; +lean_inc(x_7); +x_28 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_28, 0, x_7); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30; +lean_inc(x_7); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_7); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33; +lean_inc(x_7); +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_7); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39; +lean_inc(x_8); +lean_inc(x_9); +x_34 = l_Lean_addMacroScope(x_9, x_33, x_8); +x_35 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37; +x_36 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41; +lean_inc(x_7); +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_7); +lean_ctor_set(x_37, 1, x_35); +lean_ctor_set(x_37, 2, x_34); +lean_ctor_set(x_37, 3, x_36); +lean_inc(x_21); +lean_inc(x_7); +x_38 = l_Lean_Syntax_node1(x_7, x_14, x_21); +x_39 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35; +lean_inc(x_7); +x_40 = l_Lean_Syntax_node2(x_7, x_39, x_37, x_38); +x_41 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32; +lean_inc(x_7); +x_42 = l_Lean_Syntax_node2(x_7, x_41, x_32, x_40); +x_43 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42; +lean_inc(x_7); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_7); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29; +lean_inc(x_7); +x_46 = l_Lean_Syntax_node3(x_7, x_45, x_30, x_42, x_44); +x_47 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27; +lean_inc(x_16); +lean_inc(x_7); +x_48 = l_Lean_Syntax_node2(x_7, x_47, x_16, x_46); +x_49 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43; +lean_inc(x_7); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_7); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48; +lean_inc(x_8); +lean_inc(x_9); +x_52 = l_Lean_addMacroScope(x_9, x_51, x_8); +x_53 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47; +x_54 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51; +lean_inc(x_7); +x_55 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_55, 0, x_7); +lean_ctor_set(x_55, 1, x_53); +lean_ctor_set(x_55, 2, x_52); +lean_ctor_set(x_55, 3, x_54); +x_56 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45; +lean_inc(x_7); +x_57 = l_Lean_Syntax_node1(x_7, x_56, x_55); +x_58 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11; +lean_inc(x_16); +lean_inc(x_7); +x_59 = l_Lean_Syntax_node2(x_7, x_58, x_57, x_16); +x_60 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55; +x_61 = l_Lean_addMacroScope(x_9, x_60, x_8); +x_62 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53; +x_63 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57; +lean_inc(x_7); +x_64 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_64, 0, x_7); +lean_ctor_set(x_64, 1, x_62); +lean_ctor_set(x_64, 2, x_61); +lean_ctor_set(x_64, 3, x_63); +lean_inc(x_21); +lean_inc(x_7); +x_65 = l_Lean_Syntax_node2(x_7, x_14, x_21, x_2); +lean_inc(x_7); +x_66 = l_Lean_Syntax_node2(x_7, x_39, x_64, x_65); +lean_inc(x_7); +x_67 = l_Lean_Syntax_node1(x_7, x_56, x_66); +lean_inc(x_16); +lean_inc(x_7); +x_68 = l_Lean_Syntax_node2(x_7, x_58, x_67, x_16); +lean_inc(x_7); +x_69 = l_Lean_Syntax_node2(x_7, x_14, x_59, x_68); +x_70 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7; +lean_inc(x_7); +x_71 = l_Lean_Syntax_node1(x_7, x_70, x_69); +x_72 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24; +lean_inc_n(x_16, 2); +lean_inc(x_7); +x_73 = l_Lean_Syntax_node6(x_7, x_72, x_28, x_48, x_50, x_71, x_16, x_16); +lean_inc(x_16); +lean_inc(x_7); +x_74 = l_Lean_Syntax_node2(x_7, x_58, x_73, x_16); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_75 = l_Lean_quoteNameMk(x_25); +x_76 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18; +lean_inc_n(x_16, 2); +lean_inc(x_7); +x_77 = l_Lean_Syntax_node5(x_7, x_76, x_21, x_16, x_16, x_23, x_75); +x_78 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16; +lean_inc(x_7); +x_79 = l_Lean_Syntax_node1(x_7, x_78, x_77); +x_80 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13; +lean_inc(x_16); +lean_inc(x_7); +x_81 = l_Lean_Syntax_node3(x_7, x_80, x_13, x_16, x_79); +lean_inc(x_7); +x_82 = l_Lean_Syntax_node2(x_7, x_58, x_81, x_16); +lean_inc(x_7); +x_83 = l_Lean_Syntax_node2(x_7, x_14, x_82, x_74); +lean_inc(x_7); +x_84 = l_Lean_Syntax_node1(x_7, x_70, x_83); +x_85 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4; +x_86 = l_Lean_Syntax_node2(x_7, x_85, x_11, x_84); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_4); +return x_87; } else { -lean_object* x_52; size_t x_53; size_t x_54; -lean_inc(x_3); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_3); -lean_ctor_set(x_52, 1, x_48); -x_53 = 1; -x_54 = lean_usize_add(x_6, x_53); -x_6 = x_54; -x_7 = x_52; -goto _start; -} +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_25); +x_88 = lean_ctor_get(x_26, 0); +lean_inc(x_88); +lean_dec(x_26); +x_89 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60; +x_90 = l_String_intercalate(x_89, x_88); +x_91 = l_Lean_Meta_Grind_updateLastTag___closed__4; +x_92 = lean_string_append(x_91, x_90); +lean_dec(x_90); +x_93 = lean_box(2); +x_94 = l_Lean_Syntax_mkNameLit(x_92, x_93); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_19); +x_96 = lean_array_mk(x_95); +x_97 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59; +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_93); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_96); +x_99 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18; +lean_inc_n(x_16, 2); +lean_inc(x_7); +x_100 = l_Lean_Syntax_node5(x_7, x_99, x_21, x_16, x_16, x_23, x_98); +x_101 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16; +lean_inc(x_7); +x_102 = l_Lean_Syntax_node1(x_7, x_101, x_100); +x_103 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13; +lean_inc(x_16); +lean_inc(x_7); +x_104 = l_Lean_Syntax_node3(x_7, x_103, x_13, x_16, x_102); +lean_inc(x_7); +x_105 = l_Lean_Syntax_node2(x_7, x_58, x_104, x_16); +lean_inc(x_7); +x_106 = l_Lean_Syntax_node2(x_7, x_14, x_105, x_74); +lean_inc(x_7); +x_107 = l_Lean_Syntax_node1(x_7, x_70, x_106); +x_108 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4; +x_109 = l_Lean_Syntax_node2(x_7, x_108, x_11, x_107); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_4); +return x_110; } } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("interpolatedStrKind", 19, 19); +return x_1; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_box(0); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_7); -x_12 = lean_array_size(x_8); -x_13 = 0; -x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_8, x_9, x_10, x_8, x_12, x_13, x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = 1; -return x_16; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -return x_18; } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("typeAscription", 14, 14); +return x_1; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_dec_eq(x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_1); -x_9 = 0; -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1; +x_3 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2; +x_4 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } -else -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_box(0); -x_11 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_10); -return x_11; } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6() { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); -if (x_7 == 0) -{ -uint8_t x_8; -lean_dec(x_4); -x_8 = 0; -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("MessageData", 11, 11); +return x_1; } -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_ctor_get(x_5, 0); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_10, x_11); -lean_dec(x_11); -if (x_12 == 0) +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7() { +_start: { -lean_object* x_13; lean_object* x_14; -lean_dec(x_9); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_add(x_4, x_13); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_14; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8() { +_start: { -lean_object* x_16; uint8_t x_17; -x_16 = lean_box(0); -x_17 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_9, x_5, x_16); -if (x_17 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9() { +_start: { -lean_object* x_18; lean_object* x_19; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_4, x_18); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_19; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10() { +_start: { -uint8_t x_21; -lean_dec(x_4); -x_21 = 1; -return x_21; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); -x_8 = lean_ctor_get(x_2, 1); -x_9 = lean_box(0); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_7); -x_12 = lean_array_size(x_8); -x_13 = 0; -x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_8, x_9, x_10, x_8, x_12, x_13, x_11); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14() { +_start: { -uint8_t x_16; -x_16 = 1; -return x_16; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termM!_", 7, 7); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15() { +_start: { -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -return x_18; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__6; +x_2 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("m!", 2, 2); +return x_1; +} } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_array_get_size(x_4); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_6); -lean_dec(x_6); -x_8 = lean_nat_dec_eq(x_5, x_7); -lean_dec(x_7); -lean_dec(x_5); -if (x_8 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) { -uint8_t x_9; +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); lean_dec(x_1); -x_9 = 0; -return x_9; +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; } else { -lean_object* x_10; uint8_t x_11; -x_10 = lean_box(0); -x_11 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_10); -return x_11; -} +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(3u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +lean_inc(x_11); +x_12 = l_Lean_Syntax_getKind(x_11); +x_13 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2; +x_14 = lean_name_eq(x_12, x_13); +lean_dec(x_12); +if (x_14 == 0) +{ +uint8_t x_54; +x_54 = 0; +x_15 = x_54; +goto block_53; } +else +{ +uint8_t x_55; +x_55 = 1; +x_15 = x_55; +goto block_53; } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: +block_53: { -if (lean_obj_tag(x_1) == 0) +if (x_15 == 0) { -lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_7 = lean_usize_land(x_2, x_6); -x_8 = lean_usize_to_nat(x_7); -x_9 = lean_box(2); -x_10 = lean_array_get(x_9, x_4, x_8); -lean_dec(x_8); -lean_dec(x_4); -switch (lean_obj_tag(x_10)) { -case 0: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_ctor_get(x_3, 0); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -x_14 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) +lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_16 = lean_ctor_get(x_2, 5); +lean_inc(x_16); +x_17 = 0; +x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); +lean_dec(x_16); +x_19 = lean_ctor_get(x_2, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 1); +lean_inc(x_20); +x_21 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30; +lean_inc(x_18); +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_18); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5; +lean_inc(x_18); +x_24 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8; +x_26 = l_Lean_addMacroScope(x_20, x_25, x_19); +x_27 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7; +x_28 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13; +lean_inc(x_18); +x_29 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_29, 0, x_18); +lean_ctor_set(x_29, 1, x_27); +lean_ctor_set(x_29, 2, x_26); +lean_ctor_set(x_29, 3, x_28); +x_30 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9; +lean_inc(x_18); +x_31 = l_Lean_Syntax_node1(x_18, x_30, x_29); +x_32 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42; +lean_inc(x_18); +x_33 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_33, 0, x_18); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4; +x_35 = l_Lean_Syntax_node5(x_18, x_34, x_22, x_11, x_24, x_31, x_33); +x_36 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(x_9, x_35, x_2, x_3); +lean_dec(x_9); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -uint8_t x_15; -lean_dec(x_11); -x_15 = 0; -return x_15; +return x_36; } else { -lean_object* x_16; uint8_t x_17; -x_16 = lean_box(0); -x_17 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_11, x_3, x_16); -return x_17; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; } } -case 1: +else { -lean_object* x_18; size_t x_19; -x_18 = lean_ctor_get(x_10, 0); -lean_inc(x_18); -lean_dec(x_10); -x_19 = lean_usize_shift_right(x_2, x_5); -x_1 = x_18; -x_2 = x_19; -goto _start; +lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_41 = lean_ctor_get(x_2, 5); +lean_inc(x_41); +x_42 = 0; +x_43 = l_Lean_SourceInfo_fromRef(x_41, x_42); +lean_dec(x_41); +x_44 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16; +lean_inc(x_43); +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15; +x_47 = l_Lean_Syntax_node2(x_43, x_46, x_45, x_11); +x_48 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(x_9, x_47, x_2, x_3); +lean_dec(x_9); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) +{ +return x_48; } -default: +else { -uint8_t x_21; -x_21 = 0; -return x_21; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_48, 0); +x_51 = lean_ctor_get(x_48, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_48); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_ctor_get(x_1, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_1, 1); -lean_inc(x_23); -lean_dec(x_1); -x_24 = lean_unsigned_to_nat(0u); -x_25 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_22, x_23, lean_box(0), x_24, x_3); -lean_dec(x_23); -lean_dec(x_22); -return x_25; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint64_t x_3; lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint64_t x_8; size_t x_9; uint8_t x_10; -x_3 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_box(0); -x_6 = lean_array_size(x_4); -x_7 = 0; -x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_4, x_5, x_4, x_6, x_7, x_3); -x_9 = lean_uint64_to_usize(x_8); -x_10 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(x_1, x_9, x_2); -return x_10; +lean_object* x_5; +x_5 = l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_st_ref_take(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_16 = lean_ctor_get(x_13, 12); -x_17 = lean_box(0); -x_18 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_16, x_1, x_17); -lean_ctor_set(x_13, 12, x_18); -x_19 = lean_st_ref_set(x_3, x_13, x_14); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_21; uint8_t x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -lean_dec(x_21); -x_22 = 1; -x_23 = lean_box(x_22); -lean_ctor_set(x_19, 0, x_23); -return x_19; +return x_6; } else { -lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_24); -lean_dec(x_19); -x_25 = 1; -x_26 = lean_box(x_25); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_24); -return x_27; +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; } } -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; -x_28 = lean_ctor_get(x_13, 0); -x_29 = lean_ctor_get(x_13, 1); -x_30 = lean_ctor_get(x_13, 2); -x_31 = lean_ctor_get(x_13, 3); -x_32 = lean_ctor_get(x_13, 4); -x_33 = lean_ctor_get(x_13, 5); -x_34 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); -x_35 = lean_ctor_get(x_13, 6); -x_36 = lean_ctor_get(x_13, 7); -x_37 = lean_ctor_get(x_13, 8); -x_38 = lean_ctor_get(x_13, 9); -x_39 = lean_ctor_get(x_13, 10); -x_40 = lean_ctor_get(x_13, 11); -x_41 = lean_ctor_get(x_13, 12); -x_42 = lean_ctor_get(x_13, 13); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_13); -x_43 = lean_box(0); -x_44 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_41, x_1, x_43); -x_45 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_45, 0, x_28); -lean_ctor_set(x_45, 1, x_29); -lean_ctor_set(x_45, 2, x_30); -lean_ctor_set(x_45, 3, x_31); -lean_ctor_set(x_45, 4, x_32); -lean_ctor_set(x_45, 5, x_33); -lean_ctor_set(x_45, 6, x_35); -lean_ctor_set(x_45, 7, x_36); -lean_ctor_set(x_45, 8, x_37); -lean_ctor_set(x_45, 9, x_38); -lean_ctor_set(x_45, 10, x_39); -lean_ctor_set(x_45, 11, x_40); -lean_ctor_set(x_45, 12, x_44); -lean_ctor_set(x_45, 13, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*14, x_34); -x_46 = lean_st_ref_set(x_3, x_45, x_14); -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -if (lean_is_exclusive(x_46)) { - lean_ctor_release(x_46, 0); - lean_ctor_release(x_46, 1); - x_48 = x_46; -} else { - lean_dec_ref(x_46); - x_48 = lean_box(0); } -x_49 = 1; -x_50 = lean_box(x_49); -if (lean_is_scalar(x_48)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_48; +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; } -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_47); -return x_51; +else +{ +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_2); -x_13 = lean_st_ref_get(x_3, x_11); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_15, 12); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_17, x_12); -if (x_18 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_19; lean_object* x_20; -lean_free_object(x_13); -x_19 = lean_box(0); -x_20 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); -return x_20; +lean_dec(x_5); +return x_6; } else { -uint8_t x_21; lean_object* x_22; +lean_object* x_9; lean_object* x_10; uint64_t x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; uint64_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; size_t x_21; size_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_9); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +x_13 = lean_box(0); +x_14 = lean_array_size(x_12); +x_15 = 0; +x_16 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_12, x_13, x_12, x_14, x_15, x_11); lean_dec(x_12); -x_21 = 0; -x_22 = lean_box(x_21); -lean_ctor_set(x_13, 0, x_22); -return x_13; +x_17 = lean_uint64_to_usize(x_16); +x_18 = 1; +x_19 = lean_usize_sub(x_1, x_18); +x_20 = 5; +x_21 = lean_usize_mul(x_20, x_19); +x_22 = lean_usize_shift_right(x_17, x_21); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_5, x_23); +lean_dec(x_5); +x_25 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_6, x_22, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_24; +x_6 = x_25; +goto _start; } } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_23 = lean_ctor_get(x_13, 0); -x_24 = lean_ctor_get(x_13, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_13); -x_25 = lean_ctor_get(x_23, 12); -lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_25, x_12); -if (x_26 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) { -lean_object* x_27; lean_object* x_28; -x_27 = lean_box(0); -x_28 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); -return x_28; +lean_dec(x_3); +return x_7; } else { -uint8_t x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); lean_dec(x_12); -x_29 = 0; -x_30 = lean_box(x_29); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_24); -return x_31; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) { -size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_uint64(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_1, x_2, x_3, x_7, x_8, x_9); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_11 = lean_box_uint64(x_10); -return x_11; -} +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_uint64(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_1, x_2, x_3, x_7, x_8, x_9); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_11 = lean_box_uint64(x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) { -size_t x_7; lean_object* x_8; -x_7 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_7, x_2, x_3, x_4, x_5, x_6); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; lean_dec(x_3); -lean_dec(x_2); -return x_8; -} +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_3); +lean_object* x_34; lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; -} +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(x_7); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_1, x_6, x_7, x_4, x_5); -return x_8; +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); } +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) { -size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_uint64(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_1, x_2, x_3, x_7, x_8, x_9); +lean_object* x_50; lean_object* x_51; lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_11 = lean_box_uint64(x_10); -return x_11; -} +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { _start: { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_1, x_2, x_3); +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; -} +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -uint8_t x_6; lean_object* x_7; -x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_7 = lean_box(x_6); +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); return x_7; } -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) { -uint8_t x_4; lean_object* x_5; -x_4 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) { -size_t x_4; uint8_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(x_1, x_4, x_3); +lean_object* x_25; lean_dec(x_3); -x_6 = lean_box(x_5); -return x_6; -} +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_1, x_2); -lean_dec(x_2); -x_4 = lean_box(x_3); -return x_4; +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +if (x_33 == 0) +{ +lean_object* x_34; lean_dec(x_3); -lean_dec(x_2); -return x_12; -} +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_markTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_st_ref_take(x_4, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = !lean_is_exclusive(x_14); -if (x_16 == 0) +} +} +else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_17 = lean_ctor_get(x_14, 13); -x_18 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_18, 0, x_1); -lean_ctor_set(x_18, 1, x_2); -lean_ctor_set(x_18, 2, x_3); -x_19 = l_Std_Queue_enqueue___rarg(x_18, x_17); -lean_ctor_set(x_14, 13, x_19); -x_20 = lean_st_ref_set(x_4, x_14, x_15); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(x_7); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -x_23 = lean_box(0); -lean_ctor_set(x_20, 0, x_23); -return x_20; +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_box(0); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); +} +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; } +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_27 = lean_ctor_get(x_14, 0); -x_28 = lean_ctor_get(x_14, 1); -x_29 = lean_ctor_get(x_14, 2); -x_30 = lean_ctor_get(x_14, 3); -x_31 = lean_ctor_get(x_14, 4); -x_32 = lean_ctor_get(x_14, 5); -x_33 = lean_ctor_get_uint8(x_14, sizeof(void*)*14); -x_34 = lean_ctor_get(x_14, 6); -x_35 = lean_ctor_get(x_14, 7); -x_36 = lean_ctor_get(x_14, 8); -x_37 = lean_ctor_get(x_14, 9); -x_38 = lean_ctor_get(x_14, 10); -x_39 = lean_ctor_get(x_14, 11); -x_40 = lean_ctor_get(x_14, 12); -x_41 = lean_ctor_get(x_14, 13); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_14); -x_42 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_42, 0, x_1); -lean_ctor_set(x_42, 1, x_2); -lean_ctor_set(x_42, 2, x_3); -x_43 = l_Std_Queue_enqueue___rarg(x_42, x_41); -x_44 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_44, 0, x_27); -lean_ctor_set(x_44, 1, x_28); -lean_ctor_set(x_44, 2, x_29); -lean_ctor_set(x_44, 3, x_30); -lean_ctor_set(x_44, 4, x_31); -lean_ctor_set(x_44, 5, x_32); -lean_ctor_set(x_44, 6, x_34); -lean_ctor_set(x_44, 7, x_35); -lean_ctor_set(x_44, 8, x_36); -lean_ctor_set(x_44, 9, x_37); -lean_ctor_set(x_44, 10, x_38); -lean_ctor_set(x_44, 11, x_39); -lean_ctor_set(x_44, 12, x_40); -lean_ctor_set(x_44, 13, x_43); -lean_ctor_set_uint8(x_44, sizeof(void*)*14, x_33); -x_45 = lean_st_ref_set(x_4, x_44, x_15); -x_46 = lean_ctor_get(x_45, 1); -lean_inc(x_46); -if (lean_is_exclusive(x_45)) { - lean_ctor_release(x_45, 0); - lean_ctor_release(x_45, 1); - x_47 = x_45; -} else { - lean_dec_ref(x_45); - x_47 = lean_box(0); +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; } -x_48 = lean_box(0); -if (lean_is_scalar(x_47)) { - x_49 = lean_alloc_ctor(0, 2, 0); -} else { - x_49 = x_47; } -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_46); -return x_49; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_addNewFact(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; -} +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = 1; +return x_16; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_13 = l_Lean_Meta_Grind_addNewFact(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_st_ref_take(x_4, x_14); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_16, 11); -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_add(x_19, x_20); -lean_dec(x_19); -lean_ctor_set(x_16, 11, x_21); -x_22 = lean_st_ref_set(x_4, x_16, x_17); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_22, 0); -lean_dec(x_24); -x_25 = lean_box(0); -lean_ctor_set(x_22, 0, x_25); -return x_22; +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_22, 1); -lean_inc(x_26); -lean_dec(x_22); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_26); -return x_28; -} -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_29 = lean_ctor_get(x_16, 0); -x_30 = lean_ctor_get(x_16, 1); -x_31 = lean_ctor_get(x_16, 2); -x_32 = lean_ctor_get(x_16, 3); -x_33 = lean_ctor_get(x_16, 4); -x_34 = lean_ctor_get(x_16, 5); -x_35 = lean_ctor_get_uint8(x_16, sizeof(void*)*14); -x_36 = lean_ctor_get(x_16, 6); -x_37 = lean_ctor_get(x_16, 7); -x_38 = lean_ctor_get(x_16, 8); -x_39 = lean_ctor_get(x_16, 9); -x_40 = lean_ctor_get(x_16, 10); -x_41 = lean_ctor_get(x_16, 11); -x_42 = lean_ctor_get(x_16, 12); -x_43 = lean_ctor_get(x_16, 13); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_16); -x_44 = lean_unsigned_to_nat(1u); -x_45 = lean_nat_add(x_41, x_44); -lean_dec(x_41); -x_46 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_46, 0, x_29); -lean_ctor_set(x_46, 1, x_30); -lean_ctor_set(x_46, 2, x_31); -lean_ctor_set(x_46, 3, x_32); -lean_ctor_set(x_46, 4, x_33); -lean_ctor_set(x_46, 5, x_34); -lean_ctor_set(x_46, 6, x_36); -lean_ctor_set(x_46, 7, x_37); -lean_ctor_set(x_46, 8, x_38); -lean_ctor_set(x_46, 9, x_39); -lean_ctor_set(x_46, 10, x_40); -lean_ctor_set(x_46, 11, x_45); -lean_ctor_set(x_46, 12, x_42); -lean_ctor_set(x_46, 13, x_43); -lean_ctor_set_uint8(x_46, sizeof(void*)*14, x_35); -x_47 = lean_st_ref_set(x_4, x_46, x_17); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_47)) { - lean_ctor_release(x_47, 0); - lean_ctor_release(x_47, 1); - x_49 = x_47; -} else { - lean_dec_ref(x_47); - x_49 = lean_box(0); -} -x_50 = lean_box(0); -if (lean_is_scalar(x_49)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_49; -} -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -return x_51; +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_10); +return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +if (x_8 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_st_ref_get(x_1, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_15, 3); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_ctor_get(x_11, 11); -lean_inc(x_17); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); lean_dec(x_11); -x_18 = lean_nat_dec_le(x_16, x_17); -lean_dec(x_17); -lean_dec(x_16); -x_19 = lean_box(x_18); -lean_ctor_set(x_13, 0, x_19); -return x_13; +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -x_22 = lean_ctor_get(x_20, 3); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_11, 11); -lean_inc(x_23); -lean_dec(x_11); -x_24 = lean_nat_dec_le(x_22, x_23); -lean_dec(x_23); -lean_dec(x_22); -x_25 = lean_box(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_21); -return x_26; -} +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_ctor_get(x_3, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +x_20 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_17); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_2, x_21); lean_dec(x_2); -lean_dec(x_1); -return x_10; +x_2 = x_22; +goto _start; } +else +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_box(0); +x_25 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_17, x_3, x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_6); +lean_dec(x_5); +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_2, x_26); +lean_dec(x_2); +x_2 = x_27; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_1); +if (x_29 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_14, x_1); -lean_ctor_set(x_11, 0, x_15); -return x_11; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_1, 1); +lean_dec(x_30); +x_31 = lean_ctor_get(x_1, 0); +lean_dec(x_31); +x_32 = lean_array_fset(x_5, x_2, x_3); +x_33 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_33); +lean_ctor_set(x_1, 0, x_32); +return x_1; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_ctor_get(x_11, 0); -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_11); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_18, x_1); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -return x_20; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_1); +x_34 = lean_array_fset(x_5, x_2, x_3); +x_35 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 5); -x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_6, x_7, x_8, x_9, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_11); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_11); -lean_ctor_set(x_15, 1, x_14); -lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 0, x_15); -return x_12; +uint8_t x_16; +x_16 = 1; +return x_16; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_12, 0); -x_17 = lean_ctor_get(x_12, 1); +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_12); -lean_inc(x_11); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_11); -lean_ctor_set(x_18, 1, x_16); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; } } } -static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__1() { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("internal `grind` error, term has not been internalized", 54, 54); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__2() { -_start: +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_getENode___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; } -static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_10); +return x_11; } -static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_getENode___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_11, 1); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); lean_dec(x_13); -x_16 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_15, x_1); -if (lean_obj_tag(x_16) == 0) +if (x_14 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_free_object(x_11); -x_17 = l_Lean_indentExpr(x_1); -x_18 = l_Lean_Meta_Grind_getENode___closed__2; -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_Meta_Grind_getENode___closed__4; -x_21 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -return x_22; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; } else { -lean_object* x_23; -lean_dec(x_1); -x_23 = lean_ctor_get(x_16, 0); -lean_inc(x_23); -lean_dec(x_16); -lean_ctor_set(x_11, 0, x_23); -return x_11; -} +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_22); +lean_dec(x_22); +lean_dec(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_free_object(x_15); +x_24 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_array_fset(x_17, x_12, x_25); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_26); +return x_1; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_11, 0); -x_25 = lean_ctor_get(x_11, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_11); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_26, x_1); -if (lean_obj_tag(x_27) == 0) +uint8_t x_27; +lean_inc(x_19); +x_27 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_19, x_4, x_16); +if (x_27 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_28 = l_Lean_indentExpr(x_1); -x_29 = l_Lean_Meta_Grind_getENode___closed__2; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -x_31 = l_Lean_Meta_Grind_getENode___closed__4; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -return x_33; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_free_object(x_15); +x_28 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_array_fset(x_17, x_12, x_29); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_30); +return x_1; } else { -lean_object* x_34; lean_object* x_35; -lean_dec(x_1); -x_34 = lean_ctor_get(x_27, 0); -lean_inc(x_34); -lean_dec(x_27); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_25); -return x_35; -} +lean_object* x_31; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_31 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_31); +return x_1; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 8); -lean_inc(x_14); -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_14); -return x_11; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_11, 0); -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_11); -x_17 = lean_ctor_get(x_15, 8); -lean_inc(x_17); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_15, 0); +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_inc(x_32); lean_dec(x_15); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} +x_34 = lean_ctor_get(x_4, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +x_36 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_array_fset(x_17, x_12, x_38); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); +return x_1; } else { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) +uint8_t x_40; +lean_inc(x_32); +x_40 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_32, x_4, x_16); +if (x_40 == 0) { -return x_11; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_32, x_33, x_4, x_5); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_43 = lean_array_fset(x_17, x_12, x_42); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_43); +return x_1; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_11, 0); -x_21 = lean_ctor_get(x_11, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_11); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} +lean_object* x_44; lean_object* x_45; +lean_dec(x_33); +lean_dec(x_32); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_4); +lean_ctor_set(x_44, 1, x_5); +x_45 = lean_array_fset(x_17, x_12, x_44); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_45); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getGeneration(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +case 1: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_13); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_15); +if (x_46 == 0) { -lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_12, 2); -lean_inc(x_17); +lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_15, 0); +x_48 = lean_usize_shift_right(x_2, x_9); +x_49 = lean_usize_add(x_3, x_8); +x_50 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_47, x_48, x_49, x_4, x_5); +lean_ctor_set(x_15, 0, x_50); +x_51 = lean_array_fset(x_17, x_12, x_15); lean_dec(x_12); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_16); -lean_dec(x_16); -lean_dec(x_17); -x_19 = lean_box(x_18); -lean_ctor_set(x_14, 0, x_19); -return x_14; +lean_ctor_set(x_1, 0, x_51); +return x_1; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; -x_20 = lean_ctor_get(x_14, 0); -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_14); -x_22 = lean_ctor_get(x_12, 2); -lean_inc(x_22); +lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_52 = lean_ctor_get(x_15, 0); +lean_inc(x_52); +lean_dec(x_15); +x_53 = lean_usize_shift_right(x_2, x_9); +x_54 = lean_usize_add(x_3, x_8); +x_55 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_52, x_53, x_54, x_4, x_5); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_57 = lean_array_fset(x_17, x_12, x_56); lean_dec(x_12); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_22, x_20); -lean_dec(x_20); -lean_dec(x_22); -x_24 = lean_box(x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_21); -return x_25; -} +lean_ctor_set(x_1, 0, x_57); +return x_1; } -else -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_11); -if (x_26 == 0) -{ -return x_11; } -else +default: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_11, 0); -x_28 = lean_ctor_get(x_11, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_11); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_58; lean_object* x_59; +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_4); +lean_ctor_set(x_58, 1, x_5); +x_59 = lean_array_fset(x_17, x_12, x_58); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_59); +return x_1; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_60; size_t x_61; size_t x_62; size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_60 = lean_ctor_get(x_1, 0); +lean_inc(x_60); +lean_dec(x_1); +x_61 = 1; +x_62 = 5; +x_63 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_64 = lean_usize_land(x_2, x_63); +x_65 = lean_usize_to_nat(x_64); +x_66 = lean_array_get_size(x_60); +x_67 = lean_nat_dec_lt(x_65, x_66); +lean_dec(x_66); +if (x_67 == 0) +{ +lean_object* x_68; +lean_dec(x_65); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} +x_68 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_68, 0, x_60); +return x_68; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_13); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_array_fget(x_60, x_65); +x_70 = lean_box(0); +x_71 = lean_array_fset(x_60, x_65, x_70); +switch (lean_obj_tag(x_69)) { +case 0: { -lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_12, 2); -lean_inc(x_17); -lean_dec(x_12); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_16); -lean_dec(x_16); -lean_dec(x_17); -x_19 = lean_box(x_18); -lean_ctor_set(x_14, 0, x_19); -return x_14; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_72 = lean_ctor_get(x_69, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_69, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_74 = x_69; +} else { + lean_dec_ref(x_69); + x_74 = lean_box(0); } -else +x_75 = lean_ctor_get(x_4, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_72, 0); +lean_inc(x_76); +x_77 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_75, x_76); +lean_dec(x_76); +lean_dec(x_75); +if (x_77 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; -x_20 = lean_ctor_get(x_14, 0); -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_14); -x_22 = lean_ctor_get(x_12, 2); -lean_inc(x_22); -lean_dec(x_12); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_22, x_20); -lean_dec(x_20); -lean_dec(x_22); -x_24 = lean_box(x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_21); -return x_25; -} +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_74); +x_78 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); +x_79 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_79, 0, x_78); +x_80 = lean_array_fset(x_71, x_65, x_79); +lean_dec(x_65); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; } else { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_11); -if (x_26 == 0) +uint8_t x_82; +lean_inc(x_72); +x_82 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_72, x_4, x_70); +if (x_82 == 0) { -return x_11; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_74); +x_83 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_72, x_73, x_4, x_5); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +x_85 = lean_array_fset(x_71, x_65, x_84); +lean_dec(x_65); +x_86 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_86, 0, x_85); +return x_86; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_11, 0); -x_28 = lean_ctor_get(x_11, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_11); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_73); +lean_dec(x_72); +if (lean_is_scalar(x_74)) { + x_87 = lean_alloc_ctor(0, 2, 0); +} else { + x_87 = x_74; } +lean_ctor_set(x_87, 0, x_4); +lean_ctor_set(x_87, 1, x_5); +x_88 = lean_array_fset(x_71, x_65, x_87); +lean_dec(x_65); +x_89 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_89, 0, x_88); +return x_89; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +case 1: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_isEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; +lean_object* x_90; lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_90 = lean_ctor_get(x_69, 0); +lean_inc(x_90); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + x_91 = x_69; +} else { + lean_dec_ref(x_69); + x_91 = lean_box(0); } +x_92 = lean_usize_shift_right(x_2, x_62); +x_93 = lean_usize_add(x_3, x_61); +x_94 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_90, x_92, x_93, x_4, x_5); +if (lean_is_scalar(x_91)) { + x_95 = lean_alloc_ctor(1, 1, 0); +} else { + x_95 = x_91; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_2); -if (x_12 == 0) -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_getENode(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Meta_Grind_getENode(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -if (lean_obj_tag(x_16) == 0) -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_14, 2); -lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_ctor_get(x_18, 2); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_19, x_20); -lean_dec(x_20); -lean_dec(x_19); -x_22 = lean_box(x_21); -lean_ctor_set(x_16, 0, x_22); -return x_16; +lean_ctor_set(x_95, 0, x_94); +x_96 = lean_array_fset(x_71, x_65, x_95); +lean_dec(x_65); +x_97 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_97, 0, x_96); +return x_97; } -else +default: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_23 = lean_ctor_get(x_16, 0); -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_16); -x_25 = lean_ctor_get(x_14, 2); -lean_inc(x_25); -lean_dec(x_14); -x_26 = lean_ctor_get(x_23, 2); -lean_inc(x_26); -lean_dec(x_23); -x_27 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_25, x_26); -lean_dec(x_26); -lean_dec(x_25); -x_28 = lean_box(x_27); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_24); -return x_29; -} +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_4); +lean_ctor_set(x_98, 1, x_5); +x_99 = lean_array_fset(x_71, x_65, x_98); +lean_dec(x_65); +x_100 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_100, 0, x_99); +return x_100; } -else -{ -uint8_t x_30; -lean_dec(x_14); -x_30 = !lean_is_exclusive(x_16); -if (x_30 == 0) -{ -return x_16; } -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_16, 0); -x_32 = lean_ctor_get(x_16, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_16); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; } } } else { -uint8_t x_34; -lean_dec(x_2); -x_34 = !lean_is_exclusive(x_13); -if (x_34 == 0) +uint8_t x_101; +x_101 = !lean_is_exclusive(x_1); +if (x_101 == 0) { -return x_13; -} -else +lean_object* x_102; lean_object* x_103; size_t x_104; uint8_t x_105; +x_102 = lean_unsigned_to_nat(0u); +x_103 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_1, x_102, x_4, x_5); +x_104 = 7; +x_105 = lean_usize_dec_le(x_104, x_3); +if (x_105 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_13, 0); -x_36 = lean_ctor_get(x_13, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_13); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} -} +lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_106 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_103); +x_107 = lean_unsigned_to_nat(4u); +x_108 = lean_nat_dec_lt(x_106, x_107); +lean_dec(x_106); +if (x_108 == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_109 = lean_ctor_get(x_103, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_103, 1); +lean_inc(x_110); +lean_dec(x_103); +x_111 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_112 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_109, x_110, lean_box(0), x_102, x_111); +lean_dec(x_110); +lean_dec(x_109); +return x_112; } else { -uint8_t x_38; lean_object* x_39; lean_object* x_40; -lean_dec(x_2); -lean_dec(x_1); -x_38 = 1; -x_39 = lean_box(x_38); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_11); -return x_40; -} +return x_103; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_isEqv(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; +return x_103; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; size_t x_118; uint8_t x_119; +x_113 = lean_ctor_get(x_1, 0); +x_114 = lean_ctor_get(x_1, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_1); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +x_116 = lean_unsigned_to_nat(0u); +x_117 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7(x_115, x_116, x_4, x_5); +x_118 = 7; +x_119 = lean_usize_dec_le(x_118, x_3); +if (x_119 == 0) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +lean_object* x_120; lean_object* x_121; uint8_t x_122; +x_120 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_117); +x_121 = lean_unsigned_to_nat(4u); +x_122 = lean_nat_dec_lt(x_120, x_121); +lean_dec(x_120); +if (x_122 == 0) { -lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = 0; -x_16 = lean_box(x_15); -lean_ctor_set(x_11, 0, x_16); -return x_11; +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_123 = lean_ctor_get(x_117, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_117, 1); +lean_inc(x_124); +lean_dec(x_117); +x_125 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_126 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_3, x_123, x_124, lean_box(0), x_116, x_125); +lean_dec(x_124); +lean_dec(x_123); +return x_126; } else { -lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_dec(x_11); -x_18 = 0; -x_19 = lean_box(x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -return x_20; +return x_117; } } else { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_11); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_11, 0); -lean_dec(x_22); -x_23 = lean_ctor_get(x_12, 0); -lean_inc(x_23); -lean_dec(x_12); -x_24 = lean_ctor_get(x_23, 2); -lean_inc(x_24); -lean_dec(x_23); -x_25 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_24, x_1); -lean_dec(x_24); -x_26 = lean_box(x_25); -lean_ctor_set(x_11, 0, x_26); -return x_11; +return x_117; } -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; -x_27 = lean_ctor_get(x_11, 1); -lean_inc(x_27); -lean_dec(x_11); -x_28 = lean_ctor_get(x_12, 0); -lean_inc(x_28); -lean_dec(x_12); -x_29 = lean_ctor_get(x_28, 2); -lean_inc(x_29); -lean_dec(x_28); -x_30 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_29, x_1); -lean_dec(x_29); -x_31 = lean_box(x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_27); -return x_32; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_isRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +uint64_t x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; uint64_t x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_4 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +x_6 = lean_box(0); +x_7 = lean_array_size(x_5); +x_8 = 0; +x_9 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_5, x_6, x_5, x_7, x_8, x_4); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; +x_10 = lean_uint64_to_usize(x_9); +x_11 = 1; +x_12 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_1, x_10, x_11, x_2, x_3); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT uint64_t l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, uint64_t x_6) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_11, 0, x_15); -return x_11; +return x_6; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_box(0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; +lean_object* x_8; uint64_t x_9; uint64_t x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_3, x_5); +x_9 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__2(x_8); +lean_dec(x_8); +x_10 = lean_uint64_mix_hash(x_6, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +x_5 = x_12; +x_6 = x_10; +goto _start; } } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) { -lean_object* x_20; uint8_t x_21; -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = !lean_is_exclusive(x_12); -if (x_21 == 0) +lean_dec(x_3); +return x_7; +} +else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_12, 0); -x_23 = lean_ctor_get(x_22, 2); -lean_inc(x_23); -lean_dec(x_22); -lean_ctor_set(x_12, 0, x_23); -return x_11; +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_12, 0); -lean_inc(x_24); -lean_dec(x_12); -x_25 = lean_ctor_get(x_24, 2); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_11, 0, x_26); -return x_11; +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_3); +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; +} +else +{ +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_27 = lean_ctor_get(x_11, 1); -lean_inc(x_27); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_dec(x_11); -x_28 = lean_ctor_get(x_12, 0); -lean_inc(x_28); -if (lean_is_exclusive(x_12)) { - lean_ctor_release(x_12, 0); - x_29 = x_12; -} else { - lean_dec_ref(x_12); - x_29 = lean_box(0); -} -x_30 = lean_ctor_get(x_28, 2); -lean_inc(x_30); -lean_dec(x_28); -if (lean_is_scalar(x_29)) { - x_31 = lean_alloc_ctor(1, 1, 0); -} else { - x_31 = x_29; +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); +lean_dec(x_9); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_3); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; } -lean_ctor_set(x_31, 0, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_27); -return x_32; +else +{ +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getRoot_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 2); -lean_inc(x_14); -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_14); -return x_11; +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_11, 0); -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_11); -x_17 = lean_ctor_get(x_15, 2); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); } +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; } -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) { -return x_11; +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_11, 0); -x_21 = lean_ctor_get(x_11, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_11); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ lean_dec(x_3); -lean_dec(x_2); -return x_11; -} +return x_7; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_getENode(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +x_16 = lean_nat_dec_lt(x_14, x_15); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_3); +return x_7; } else { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_11); -if (x_15 == 0) +uint8_t x_17; +x_17 = !lean_is_exclusive(x_11); +if (x_17 == 0) { -return x_11; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_18 = lean_ctor_get(x_11, 2); +lean_dec(x_18); +x_19 = lean_ctor_get(x_11, 1); +lean_dec(x_19); +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_array_fget(x_13, x_14); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_14, x_22); +lean_dec(x_14); +lean_ctor_set(x_11, 1, x_23); +x_24 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_21); +lean_dec(x_21); +lean_dec(x_9); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_3); +x_25 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 0, x_25); +return x_7; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 0); -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_11); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; -} -} +size_t x_26; size_t x_27; +lean_inc(x_3); +lean_ctor_set(x_7, 0, x_3); +x_26 = 1; +x_27 = lean_usize_add(x_6, x_26); +x_6 = x_27; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getRootENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_11); +x_29 = lean_array_fget(x_13, x_14); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_14, x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_15); +x_33 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_29); +lean_dec(x_29); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +if (x_33 == 0) +{ +lean_object* x_34; lean_dec(x_3); -lean_dec(x_2); -return x_11; -} +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_34); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +size_t x_35; size_t x_36; +lean_inc(x_3); +lean_ctor_set(x_7, 1, x_32); +lean_ctor_set(x_7, 0, x_3); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_6 = x_36; +goto _start; +} +} +} +} +else { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_7, 1); +lean_inc(x_38); +lean_dec(x_7); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +x_42 = lean_nat_dec_lt(x_40, x_41); +if (x_42 == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -lean_ctor_set(x_11, 0, x_14); -return x_11; +lean_object* x_43; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_9); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_38); +return x_43; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_11, 0); -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_11); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + lean_ctor_release(x_38, 2); + x_44 = x_38; +} else { + lean_dec_ref(x_38); + x_44 = lean_box(0); } +x_45 = lean_array_fget(x_39, x_40); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_add(x_40, x_46); +lean_dec(x_40); +if (lean_is_scalar(x_44)) { + x_48 = lean_alloc_ctor(0, 3, 0); +} else { + x_48 = x_44; } -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_41); +x_49 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_9, x_45); +lean_dec(x_45); +lean_dec(x_9); +if (x_49 == 0) { -return x_11; +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1; +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_11, 0); -x_21 = lean_ctor_get(x_11, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_11); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_52; size_t x_53; size_t x_54; +lean_inc(x_3); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_3); +lean_ctor_set(x_52, 1, x_48); +x_53 = 1; +x_54 = lean_usize_add(x_6, x_53); +x_6 = x_54; +x_7 = x_52; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getNext(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_10); return x_11; } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; uint8_t x_7; @@ -9908,31 +10813,120 @@ return x_8; } else { -lean_object* x_9; uint8_t x_10; +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_9 = lean_array_fget(x_1, x_4); -x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_10, x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_dec(x_9); -if (x_10 == 0) +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_4, x_13); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_14; +goto _start; +} +else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); +lean_object* x_16; uint8_t x_17; +x_16 = lean_box(0); +x_17 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_9, x_5, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_4, x_18); lean_dec(x_4); x_3 = lean_box(0); -x_4 = x_12; +x_4 = x_19; goto _start; } else { -uint8_t x_14; +uint8_t x_21; lean_dec(x_4); -x_14 = 1; -return x_14; +x_21 = 1; +return x_21; } } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_toSubarray___rarg(x_4, x_6, x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_box(0); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_7); +x_12 = lean_array_size(x_8); +x_13 = 0; +x_14 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_8, x_9, x_10, x_8, x_12, x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +return x_18; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_array_get_size(x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_6); +lean_dec(x_6); +x_8 = lean_nat_dec_eq(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_box(0); +x_11 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_10); +return x_11; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -9952,1953 +10946,5711 @@ lean_dec(x_4); switch (lean_obj_tag(x_10)) { case 0: { -lean_object* x_11; uint8_t x_12; +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); -x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); +x_12 = lean_ctor_get(x_3, 0); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +uint8_t x_15; lean_dec(x_11); -return x_12; +x_15 = 0; +return x_15; +} +else +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_box(0); +x_17 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_11, x_3, x_16); +return x_17; +} } case 1: { -lean_object* x_13; size_t x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); +lean_object* x_18; size_t x_19; +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); lean_dec(x_10); -x_14 = lean_usize_shift_right(x_2, x_5); -x_1 = x_13; -x_2 = x_14; +x_19 = lean_usize_shift_right(x_2, x_5); +x_1 = x_18; +x_2 = x_19; goto _start; } default: { -uint8_t x_16; -x_16 = 0; -return x_16; +uint8_t x_21; +x_21 = 0; +return x_21; } } } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 1); -lean_inc(x_18); +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_1, 1); +lean_inc(x_23); lean_dec(x_1); -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_17, x_18, lean_box(0), x_19, x_3); -lean_dec(x_18); -lean_dec(x_17); -return x_20; +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_22, x_23, lean_box(0), x_24, x_3); +lean_dec(x_23); +lean_dec(x_22); +return x_25; } } } -LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(lean_object* x_1, lean_object* x_2) { _start: { -uint64_t x_3; size_t x_4; uint8_t x_5; -x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_2); -return x_5; +uint64_t x_3; lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; uint64_t x_8; size_t x_9; uint8_t x_10; +x_3 = l_Lean_Meta_Grind_instHashablePreInstance_unsafe__1(x_2); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_box(0); +x_6 = lean_array_size(x_4); +x_7 = 0; +x_8 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_4, x_5, x_4, x_6, x_7, x_3); +x_9 = lean_uint64_to_usize(x_8); +x_10 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(x_1, x_9, x_2); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_14, x_1); -x_16 = lean_box(x_15); -lean_ctor_set(x_11, 0, x_16); -return x_11; +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_13, 13); +x_17 = lean_box(0); +x_18 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_16, x_1, x_17); +lean_ctor_set(x_13, 13, x_18); +x_19 = lean_st_ref_set(x_3, x_13, x_14); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 0, x_23); +return x_19; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; 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; +x_28 = lean_ctor_get(x_13, 0); +x_29 = lean_ctor_get(x_13, 1); +x_30 = lean_ctor_get(x_13, 2); +x_31 = lean_ctor_get(x_13, 3); +x_32 = lean_ctor_get(x_13, 4); +x_33 = lean_ctor_get(x_13, 5); +x_34 = lean_ctor_get_uint8(x_13, sizeof(void*)*19); +x_35 = lean_ctor_get(x_13, 6); +x_36 = lean_ctor_get(x_13, 7); +x_37 = lean_ctor_get(x_13, 8); +x_38 = lean_ctor_get(x_13, 9); +x_39 = lean_ctor_get(x_13, 10); +x_40 = lean_ctor_get(x_13, 11); +x_41 = lean_ctor_get(x_13, 12); +x_42 = lean_ctor_get(x_13, 13); +x_43 = lean_ctor_get(x_13, 14); +x_44 = lean_ctor_get(x_13, 15); +x_45 = lean_ctor_get(x_13, 16); +x_46 = lean_ctor_get(x_13, 17); +x_47 = lean_ctor_get(x_13, 18); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_13); +x_48 = lean_box(0); +x_49 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markTheoremInstance___spec__1(x_42, x_1, x_48); +x_50 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_50, 0, x_28); +lean_ctor_set(x_50, 1, x_29); +lean_ctor_set(x_50, 2, x_30); +lean_ctor_set(x_50, 3, x_31); +lean_ctor_set(x_50, 4, x_32); +lean_ctor_set(x_50, 5, x_33); +lean_ctor_set(x_50, 6, x_35); +lean_ctor_set(x_50, 7, x_36); +lean_ctor_set(x_50, 8, x_37); +lean_ctor_set(x_50, 9, x_38); +lean_ctor_set(x_50, 10, x_39); +lean_ctor_set(x_50, 11, x_40); +lean_ctor_set(x_50, 12, x_41); +lean_ctor_set(x_50, 13, x_49); +lean_ctor_set(x_50, 14, x_43); +lean_ctor_set(x_50, 15, x_44); +lean_ctor_set(x_50, 16, x_45); +lean_ctor_set(x_50, 17, x_46); +lean_ctor_set(x_50, 18, x_47); +lean_ctor_set_uint8(x_50, sizeof(void*)*19, x_34); +x_51 = lean_st_ref_set(x_3, x_50, x_14); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_53 = x_51; +} else { + lean_dec_ref(x_51); + x_53 = lean_box(0); +} +x_54 = 1; +x_55 = lean_box(x_54); +if (lean_is_scalar(x_53)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_53; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_52); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_2); +x_13 = lean_st_ref_get(x_3, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_15, 13); lean_inc(x_17); -lean_dec(x_11); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_19, x_1); -x_21 = lean_box(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_18); -return x_22; +lean_dec(x_15); +x_18 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_17, x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_free_object(x_13); +x_19 = lean_box(0); +x_20 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_20; } +else +{ +uint8_t x_21; lean_object* x_22; +lean_dec(x_12); +x_21 = 0; +x_22 = lean_box(x_21); +lean_ctor_set(x_13, 0, x_22); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_13, 0); +x_24 = lean_ctor_get(x_13, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_13); +x_25 = lean_ctor_get(x_23, 13); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_25, x_12); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_box(0); +x_28 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_12, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); +return x_28; +} +else +{ +uint8_t x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_12); +x_29 = 0; +x_30 = lean_box(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_24); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_6; lean_object* x_7; -x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_1, x_2, x_3, x_4, x_5); +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__2(x_1, x_2, x_3, x_7, x_8, x_9); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; +x_11 = lean_box_uint64(x_10); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -size_t x_4; uint8_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_3); +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__5(x_1, x_2, x_3, x_7, x_8, x_9); lean_dec(x_3); -x_6 = lean_box(x_5); -return x_6; +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box_uint64(x_10); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_1, x_2); +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markTheoremInstance___spec__4(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); lean_dec(x_2); -x_4 = lean_box(x_3); -return x_4; +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_alreadyInternalized(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__6(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__8(x_1, x_2, x_3, x_4, x_8, x_9, x_7); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_11; +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_11, 0, x_15); -return x_11; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_box(0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markTheoremInstance___spec__7___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } } -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_11, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_12, 0); -lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_ctor_get(x_21, 4); -lean_inc(x_22); -lean_dec(x_21); -lean_ctor_set(x_11, 0, x_22); -return x_11; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_11, 1); -lean_inc(x_23); -lean_dec(x_11); -x_24 = lean_ctor_get(x_12, 0); -lean_inc(x_24); -lean_dec(x_12); -x_25 = lean_ctor_get(x_24, 4); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_23); -return x_26; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markTheoremInstance___spec__3(x_1, x_6, x_7, x_4, x_5); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getTarget_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_7 = lean_unbox_usize(x_4); lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__10(x_1, x_2, x_3, x_7, x_8, x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); +x_11 = lean_box_uint64(x_10); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_st_ref_take(x_5, x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_18 = lean_ctor_get(x_15, 5); -x_19 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_2); -lean_ctor_set(x_19, 2, x_3); -lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_4); -x_20 = lean_array_push(x_18, x_19); -lean_ctor_set(x_15, 5, x_20); -x_21 = lean_st_ref_set(x_5, x_15, x_16); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -x_24 = lean_box(0); -lean_ctor_set(x_21, 0, x_24); -return x_21; +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__12(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_markTheoremInstance___spec__14(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_28 = lean_ctor_get(x_15, 0); -x_29 = lean_ctor_get(x_15, 1); -x_30 = lean_ctor_get(x_15, 2); -x_31 = lean_ctor_get(x_15, 3); -x_32 = lean_ctor_get(x_15, 4); -x_33 = lean_ctor_get(x_15, 5); -x_34 = lean_ctor_get_uint8(x_15, sizeof(void*)*14); -x_35 = lean_ctor_get(x_15, 6); -x_36 = lean_ctor_get(x_15, 7); -x_37 = lean_ctor_get(x_15, 8); -x_38 = lean_ctor_get(x_15, 9); -x_39 = lean_ctor_get(x_15, 10); -x_40 = lean_ctor_get(x_15, 11); -x_41 = lean_ctor_get(x_15, 12); -x_42 = lean_ctor_get(x_15, 13); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_15); -x_43 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_2); -lean_ctor_set(x_43, 2, x_3); -lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_4); -x_44 = lean_array_push(x_33, x_43); -x_45 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_45, 0, x_28); -lean_ctor_set(x_45, 1, x_29); -lean_ctor_set(x_45, 2, x_30); -lean_ctor_set(x_45, 3, x_31); -lean_ctor_set(x_45, 4, x_32); -lean_ctor_set(x_45, 5, x_44); -lean_ctor_set(x_45, 6, x_35); -lean_ctor_set(x_45, 7, x_36); -lean_ctor_set(x_45, 8, x_37); -lean_ctor_set(x_45, 9, x_38); -lean_ctor_set(x_45, 10, x_39); -lean_ctor_set(x_45, 11, x_40); -lean_ctor_set(x_45, 12, x_41); -lean_ctor_set(x_45, 13, x_42); -lean_ctor_set_uint8(x_45, sizeof(void*)*14, x_34); -x_46 = lean_st_ref_set(x_5, x_45, x_16); -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -if (lean_is_exclusive(x_46)) { - lean_ctor_release(x_46, 0); - lean_ctor_release(x_46, 1); - x_48 = x_46; -} else { - lean_dec_ref(x_46); - x_48 = lean_box(0); -} -x_49 = lean_box(0); -if (lean_is_scalar(x_48)) { - x_50 = lean_alloc_ctor(0, 2, 0); -} else { - x_50 = x_48; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_47); -return x_50; } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_4); -lean_dec(x_4); -x_15 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_markTheoremInstance___spec__13(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); -return x_15; +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; } } -static uint64_t _init_l_Lean_Meta_Grind_hasSameType___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_1; uint64_t x_2; -x_1 = 1; -x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); -return x_2; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_3); -if (x_8 == 0) -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_3, 0); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -uint64_t x_11; uint8_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; lean_object* x_18; -x_11 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); -x_12 = 1; -lean_ctor_set_uint8(x_9, 9, x_12); -x_13 = 2; -x_14 = lean_uint64_shift_right(x_11, x_13); -x_15 = lean_uint64_shift_left(x_14, x_13); -x_16 = l_Lean_Meta_Grind_hasSameType___closed__1; -x_17 = lean_uint64_lor(x_15, x_16); -lean_ctor_set_uint64(x_3, sizeof(void*)*7, x_17); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_18 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_21 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_20); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Lean_Meta_isExprDefEq(x_19, x_22, x_3, x_4, x_5, x_6, x_23); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -return x_24; +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +x_5 = lean_box(x_4); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_24, 0); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_24); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +size_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_markTheoremInstance___spec__11(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; } } -else -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_24); -if (x_29 == 0) +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -return x_24; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_markTheoremInstance___spec__9(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; } -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_24, 0); -x_31 = lean_ctor_get(x_24, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_24); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_markTheoremInstance___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_33; -lean_dec(x_19); -lean_dec(x_3); +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_markTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_33 = !lean_is_exclusive(x_21); -if (x_33 == 0) +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -return x_21; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_st_ref_take(x_4, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_14, 14); +x_18 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_18, 0, x_1); +lean_ctor_set(x_18, 1, x_2); +lean_ctor_set(x_18, 2, x_3); +x_19 = l_Std_Queue_enqueue___rarg(x_18, x_17); +lean_ctor_set(x_14, 14, x_19); +x_20 = lean_st_ref_set(x_4, x_14, x_15); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +x_23 = lean_box(0); +lean_ctor_set(x_20, 0, x_23); +return x_20; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_21, 0); -x_35 = lean_ctor_get(x_21, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_box(0); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +return x_26; +} +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_27 = lean_ctor_get(x_14, 0); +x_28 = lean_ctor_get(x_14, 1); +x_29 = lean_ctor_get(x_14, 2); +x_30 = lean_ctor_get(x_14, 3); +x_31 = lean_ctor_get(x_14, 4); +x_32 = lean_ctor_get(x_14, 5); +x_33 = lean_ctor_get_uint8(x_14, sizeof(void*)*19); +x_34 = lean_ctor_get(x_14, 6); +x_35 = lean_ctor_get(x_14, 7); +x_36 = lean_ctor_get(x_14, 8); +x_37 = lean_ctor_get(x_14, 9); +x_38 = lean_ctor_get(x_14, 10); +x_39 = lean_ctor_get(x_14, 11); +x_40 = lean_ctor_get(x_14, 12); +x_41 = lean_ctor_get(x_14, 13); +x_42 = lean_ctor_get(x_14, 14); +x_43 = lean_ctor_get(x_14, 15); +x_44 = lean_ctor_get(x_14, 16); +x_45 = lean_ctor_get(x_14, 17); +x_46 = lean_ctor_get(x_14, 18); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); -lean_dec(x_21); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_14); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_1); +lean_ctor_set(x_47, 1, x_2); +lean_ctor_set(x_47, 2, x_3); +x_48 = l_Std_Queue_enqueue___rarg(x_47, x_42); +x_49 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_49, 0, x_27); +lean_ctor_set(x_49, 1, x_28); +lean_ctor_set(x_49, 2, x_29); +lean_ctor_set(x_49, 3, x_30); +lean_ctor_set(x_49, 4, x_31); +lean_ctor_set(x_49, 5, x_32); +lean_ctor_set(x_49, 6, x_34); +lean_ctor_set(x_49, 7, x_35); +lean_ctor_set(x_49, 8, x_36); +lean_ctor_set(x_49, 9, x_37); +lean_ctor_set(x_49, 10, x_38); +lean_ctor_set(x_49, 11, x_39); +lean_ctor_set(x_49, 12, x_40); +lean_ctor_set(x_49, 13, x_41); +lean_ctor_set(x_49, 14, x_48); +lean_ctor_set(x_49, 15, x_43); +lean_ctor_set(x_49, 16, x_44); +lean_ctor_set(x_49, 17, x_45); +lean_ctor_set(x_49, 18, x_46); +lean_ctor_set_uint8(x_49, sizeof(void*)*19, x_33); +x_50 = lean_st_ref_set(x_4, x_49, x_15); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_52 = x_50; +} else { + lean_dec_ref(x_50); + x_52 = lean_box(0); +} +x_53 = lean_box(0); +if (lean_is_scalar(x_52)) { + x_54 = lean_alloc_ctor(0, 2, 0); +} else { + x_54 = x_52; +} +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_51); +return x_54; } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addNewFact___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_37; -lean_dec(x_3); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_addNewFact(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -x_37 = !lean_is_exclusive(x_18); -if (x_37 == 0) -{ -return x_18; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_18, 0); -x_39 = lean_ctor_get(x_18, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_18); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} +return x_13; } } -else -{ -uint64_t x_41; uint8_t x_42; uint8_t x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; uint8_t x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; uint8_t x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; lean_object* x_65; -x_41 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); -x_42 = lean_ctor_get_uint8(x_9, 0); -x_43 = lean_ctor_get_uint8(x_9, 1); -x_44 = lean_ctor_get_uint8(x_9, 2); -x_45 = lean_ctor_get_uint8(x_9, 3); -x_46 = lean_ctor_get_uint8(x_9, 4); -x_47 = lean_ctor_get_uint8(x_9, 5); -x_48 = lean_ctor_get_uint8(x_9, 6); -x_49 = lean_ctor_get_uint8(x_9, 7); -x_50 = lean_ctor_get_uint8(x_9, 8); -x_51 = lean_ctor_get_uint8(x_9, 10); -x_52 = lean_ctor_get_uint8(x_9, 11); -x_53 = lean_ctor_get_uint8(x_9, 12); -x_54 = lean_ctor_get_uint8(x_9, 13); -x_55 = lean_ctor_get_uint8(x_9, 14); -x_56 = lean_ctor_get_uint8(x_9, 15); -x_57 = lean_ctor_get_uint8(x_9, 16); -lean_dec(x_9); -x_58 = 1; -x_59 = lean_alloc_ctor(0, 0, 17); -lean_ctor_set_uint8(x_59, 0, x_42); -lean_ctor_set_uint8(x_59, 1, x_43); -lean_ctor_set_uint8(x_59, 2, x_44); -lean_ctor_set_uint8(x_59, 3, x_45); -lean_ctor_set_uint8(x_59, 4, x_46); -lean_ctor_set_uint8(x_59, 5, x_47); -lean_ctor_set_uint8(x_59, 6, x_48); -lean_ctor_set_uint8(x_59, 7, x_49); -lean_ctor_set_uint8(x_59, 8, x_50); -lean_ctor_set_uint8(x_59, 9, x_58); -lean_ctor_set_uint8(x_59, 10, x_51); -lean_ctor_set_uint8(x_59, 11, x_52); -lean_ctor_set_uint8(x_59, 12, x_53); -lean_ctor_set_uint8(x_59, 13, x_54); -lean_ctor_set_uint8(x_59, 14, x_55); -lean_ctor_set_uint8(x_59, 15, x_56); -lean_ctor_set_uint8(x_59, 16, x_57); -x_60 = 2; -x_61 = lean_uint64_shift_right(x_41, x_60); -x_62 = lean_uint64_shift_left(x_61, x_60); -x_63 = l_Lean_Meta_Grind_hasSameType___closed__1; -x_64 = lean_uint64_lor(x_62, x_63); -lean_ctor_set(x_3, 0, x_59); -lean_ctor_set_uint64(x_3, sizeof(void*)*7, x_64); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_65 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_65) == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_68 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_67); -if (lean_obj_tag(x_68) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_13 = l_Lean_Meta_Grind_addNewFact(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_st_ref_take(x_4, x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Lean_Meta_isExprDefEq(x_66, x_69, x_3, x_4, x_5, x_6, x_70); -if (lean_obj_tag(x_71) == 0) +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_16, 11); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_add(x_19, x_20); +lean_dec(x_19); +lean_ctor_set(x_16, 11, x_21); +x_22 = lean_st_ref_set(x_4, x_16, x_17); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_74 = x_71; -} else { - lean_dec_ref(x_71); - x_74 = lean_box(0); -} -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(0, 2, 0); -} else { - x_75 = x_74; -} -lean_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_22, 0, x_25); +return x_22; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_71, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_71, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_78 = x_71; -} else { - lean_dec_ref(x_71); - x_78 = lean_box(0); -} -if (lean_is_scalar(x_78)) { - x_79 = lean_alloc_ctor(1, 2, 0); -} else { - x_79 = x_78; -} -lean_ctor_set(x_79, 0, x_76); -lean_ctor_set(x_79, 1, x_77); -return x_79; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; } } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_66); -lean_dec(x_3); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_80 = lean_ctor_get(x_68, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_68, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_82 = x_68; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_29 = lean_ctor_get(x_16, 0); +x_30 = lean_ctor_get(x_16, 1); +x_31 = lean_ctor_get(x_16, 2); +x_32 = lean_ctor_get(x_16, 3); +x_33 = lean_ctor_get(x_16, 4); +x_34 = lean_ctor_get(x_16, 5); +x_35 = lean_ctor_get_uint8(x_16, sizeof(void*)*19); +x_36 = lean_ctor_get(x_16, 6); +x_37 = lean_ctor_get(x_16, 7); +x_38 = lean_ctor_get(x_16, 8); +x_39 = lean_ctor_get(x_16, 9); +x_40 = lean_ctor_get(x_16, 10); +x_41 = lean_ctor_get(x_16, 11); +x_42 = lean_ctor_get(x_16, 12); +x_43 = lean_ctor_get(x_16, 13); +x_44 = lean_ctor_get(x_16, 14); +x_45 = lean_ctor_get(x_16, 15); +x_46 = lean_ctor_get(x_16, 16); +x_47 = lean_ctor_get(x_16, 17); +x_48 = lean_ctor_get(x_16, 18); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_16); +x_49 = lean_unsigned_to_nat(1u); +x_50 = lean_nat_add(x_41, x_49); +lean_dec(x_41); +x_51 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_51, 0, x_29); +lean_ctor_set(x_51, 1, x_30); +lean_ctor_set(x_51, 2, x_31); +lean_ctor_set(x_51, 3, x_32); +lean_ctor_set(x_51, 4, x_33); +lean_ctor_set(x_51, 5, x_34); +lean_ctor_set(x_51, 6, x_36); +lean_ctor_set(x_51, 7, x_37); +lean_ctor_set(x_51, 8, x_38); +lean_ctor_set(x_51, 9, x_39); +lean_ctor_set(x_51, 10, x_40); +lean_ctor_set(x_51, 11, x_50); +lean_ctor_set(x_51, 12, x_42); +lean_ctor_set(x_51, 13, x_43); +lean_ctor_set(x_51, 14, x_44); +lean_ctor_set(x_51, 15, x_45); +lean_ctor_set(x_51, 16, x_46); +lean_ctor_set(x_51, 17, x_47); +lean_ctor_set(x_51, 18, x_48); +lean_ctor_set_uint8(x_51, sizeof(void*)*19, x_35); +x_52 = lean_st_ref_set(x_4, x_51, x_17); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_54 = x_52; } else { - lean_dec_ref(x_68); - x_82 = lean_box(0); + lean_dec_ref(x_52); + x_54 = lean_box(0); } -if (lean_is_scalar(x_82)) { - x_83 = lean_alloc_ctor(1, 2, 0); +x_55 = lean_box(0); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); } else { - x_83 = x_82; + x_56 = x_54; } -lean_ctor_set(x_83, 0, x_80); -lean_ctor_set(x_83, 1, x_81); -return x_83; +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +return x_56; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_addTheoremInstance___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -lean_dec(x_3); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_addTheoremInstance(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -x_84 = lean_ctor_get(x_65, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_65, 1); -lean_inc(x_85); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_86 = x_65; -} else { - lean_dec_ref(x_65); - x_86 = lean_box(0); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 3); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_11, 11); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_nat_dec_le(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_20, 3); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_11, 11); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_nat_dec_le(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxInstancesExceeded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_checkMaxInstancesExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_11, 17); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_nat_dec_le(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_11, 17); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_nat_dec_le(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxCaseSplit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_checkMaxCaseSplit(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Meta_Grind_getConfig___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_11, 12); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_nat_dec_le(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +x_19 = lean_box(x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_11, 12); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_nat_dec_le(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_checkMaxEmatchExceeded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_checkMaxEmatchExceeded(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_14, x_1); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_11, 0); +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_11); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_18, x_1); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_8, 5); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_6, x_7, x_8, x_9, x_10); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_11); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set_tag(x_12, 1); +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_12); +lean_inc(x_11); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_16); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +} +static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("internal `grind` error, term has not been internalized", 54, 54); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_getENode___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_getENode___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_15, x_1); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_11); +x_17 = l_Lean_indentExpr(x_1); +x_18 = l_Lean_Meta_Grind_getENode___closed__2; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_Meta_Grind_getENode___closed__3; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +return x_22; +} +else +{ +lean_object* x_23; +lean_dec(x_1); +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +lean_dec(x_16); +lean_ctor_set(x_11, 0, x_23); +return x_11; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_11, 0); +x_25 = lean_ctor_get(x_11, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_11); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_hashRoot___spec__1(x_26, x_1); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_28 = l_Lean_indentExpr(x_1); +x_29 = l_Lean_Meta_Grind_getENode___closed__2; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l_Lean_Meta_Grind_getENode___closed__3; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); +return x_33; +} +else +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_1); +x_34 = lean_ctor_get(x_27, 0); +lean_inc(x_34); +lean_dec(x_27); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_25); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_throwError___at_Lean_Meta_Grind_getENode___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 8); +lean_inc(x_14); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_14); +return x_11; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_11, 0); +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_11); +x_17 = lean_ctor_get(x_15, 8); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getGeneration___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getGeneration(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_12, 2); +lean_inc(x_17); +lean_dec(x_12); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_16); +lean_dec(x_16); +lean_dec(x_17); +x_19 = lean_box(x_18); +lean_ctor_set(x_14, 0, x_19); +return x_14; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_20 = lean_ctor_get(x_14, 0); +x_21 = lean_ctor_get(x_14, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_14); +x_22 = lean_ctor_get(x_12, 2); +lean_inc(x_22); +lean_dec(x_12); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_22, x_20); +lean_dec(x_20); +lean_dec(x_22); +x_24 = lean_box(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +return x_25; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_11); +if (x_26 == 0) +{ +return x_11; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_11, 0); +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_11); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqTrue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_12, 2); +lean_inc(x_17); +lean_dec(x_12); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_17, x_16); +lean_dec(x_16); +lean_dec(x_17); +x_19 = lean_box(x_18); +lean_ctor_set(x_14, 0, x_19); +return x_14; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_20 = lean_ctor_get(x_14, 0); +x_21 = lean_ctor_get(x_14, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_14); +x_22 = lean_ctor_get(x_12, 2); +lean_inc(x_22); +lean_dec(x_12); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_22, x_20); +lean_dec(x_20); +lean_dec(x_22); +x_24 = lean_box(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +return x_25; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_11); +if (x_26 == 0) +{ +return x_11; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_11, 0); +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_11); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_2); +if (x_12 == 0) +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_getENode(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_Grind_getENode(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_14, 2); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_ctor_get(x_18, 2); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +x_22 = lean_box(x_21); +lean_ctor_set(x_16, 0, x_22); +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_ctor_get(x_14, 2); +lean_inc(x_25); +lean_dec(x_14); +x_26 = lean_ctor_get(x_23, 2); +lean_inc(x_26); +lean_dec(x_23); +x_27 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_25, x_26); +lean_dec(x_26); +lean_dec(x_25); +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_24); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_14); +x_30 = !lean_is_exclusive(x_16); +if (x_30 == 0) +{ +return x_16; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_16, 0); +x_32 = lean_ctor_get(x_16, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_16); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +else +{ +uint8_t x_34; +lean_dec(x_2); +x_34 = !lean_is_exclusive(x_13); +if (x_34 == 0) +{ +return x_13; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_13, 0); +x_36 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_13); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +uint8_t x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_2); +lean_dec(x_1); +x_38 = 1; +x_39 = lean_box(x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_11); +return x_40; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isEqv___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_isEqv(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = 0; +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = 0; +x_19 = lean_box(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_11); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_11, 0); +lean_dec(x_22); +x_23 = lean_ctor_get(x_12, 0); +lean_inc(x_23); +lean_dec(x_12); +x_24 = lean_ctor_get(x_23, 2); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_24, x_1); +lean_dec(x_24); +x_26 = lean_box(x_25); +lean_ctor_set(x_11, 0, x_26); +return x_11; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; +x_27 = lean_ctor_get(x_11, 1); +lean_inc(x_27); +lean_dec(x_11); +x_28 = lean_ctor_get(x_12, 0); +lean_inc(x_28); +lean_dec(x_12); +x_29 = lean_ctor_get(x_28, 2); +lean_inc(x_29); +lean_dec(x_28); +x_30 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_29, x_1); +lean_dec(x_29); +x_31 = lean_box(x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_27); +return x_32; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = !lean_is_exclusive(x_12); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_22, 2); +lean_inc(x_23); +lean_dec(x_22); +lean_ctor_set(x_12, 0, x_23); +return x_11; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_12, 0); +lean_inc(x_24); +lean_dec(x_12); +x_25 = lean_ctor_get(x_24, 2); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_11, 0, x_26); +return x_11; +} +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_27 = lean_ctor_get(x_11, 1); +lean_inc(x_27); +lean_dec(x_11); +x_28 = lean_ctor_get(x_12, 0); +lean_inc(x_28); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + x_29 = x_12; +} else { + lean_dec_ref(x_12); + x_29 = lean_box(0); +} +x_30 = lean_ctor_get(x_28, 2); +lean_inc(x_30); +lean_dec(x_28); +if (lean_is_scalar(x_29)) { + x_31 = lean_alloc_ctor(1, 1, 0); +} else { + x_31 = x_29; +} +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_27); +return x_32; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getRoot_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 2); +lean_inc(x_14); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_14); +return x_11; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_11, 0); +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_11); +x_17 = lean_ctor_get(x_15, 2); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Meta_Grind_getENode(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_14; +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_11); +if (x_15 == 0) +{ +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 0); +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_11); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getRootENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getRootENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_14); +return x_11; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_11, 0); +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_11); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +return x_11; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getNext___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getNext(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; +} +else +{ +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); +lean_dec(x_11); +return x_12; +} +case 1: +{ +lean_object* x_13; size_t x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_usize_shift_right(x_2, x_5); +x_1 = x_13; +x_2 = x_14; +goto _start; +} +default: +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 1); +lean_inc(x_18); +lean_dec(x_1); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_17, x_18, lean_box(0), x_19, x_3); +lean_dec(x_18); +lean_dec(x_17); +return x_20; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_19, x_1); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_alreadyInternalized___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_alreadyInternalized___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_alreadyInternalized___spec__1(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_alreadyInternalized___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_alreadyInternalized(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Meta_Grind_getENode_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +lean_dec(x_20); +x_21 = lean_ctor_get(x_12, 0); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_ctor_get(x_21, 4); +lean_inc(x_22); +lean_dec(x_21); +lean_ctor_set(x_11, 0, x_22); +return x_11; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_11, 1); +lean_inc(x_23); +lean_dec(x_11); +x_24 = lean_ctor_get(x_12, 0); +lean_inc(x_24); +lean_dec(x_12); +x_25 = lean_ctor_get(x_24, 4); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_23); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getTarget_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getTarget_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_st_ref_take(x_5, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_15, 5); +x_19 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_2); +lean_ctor_set(x_19, 2, x_3); +lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_4); +x_20 = lean_array_push(x_18, x_19); +lean_ctor_set(x_15, 5, x_20); +x_21 = lean_st_ref_set(x_5, x_15, x_16); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_21, 0, x_24); +return x_21; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_28 = lean_ctor_get(x_15, 0); +x_29 = lean_ctor_get(x_15, 1); +x_30 = lean_ctor_get(x_15, 2); +x_31 = lean_ctor_get(x_15, 3); +x_32 = lean_ctor_get(x_15, 4); +x_33 = lean_ctor_get(x_15, 5); +x_34 = lean_ctor_get_uint8(x_15, sizeof(void*)*19); +x_35 = lean_ctor_get(x_15, 6); +x_36 = lean_ctor_get(x_15, 7); +x_37 = lean_ctor_get(x_15, 8); +x_38 = lean_ctor_get(x_15, 9); +x_39 = lean_ctor_get(x_15, 10); +x_40 = lean_ctor_get(x_15, 11); +x_41 = lean_ctor_get(x_15, 12); +x_42 = lean_ctor_get(x_15, 13); +x_43 = lean_ctor_get(x_15, 14); +x_44 = lean_ctor_get(x_15, 15); +x_45 = lean_ctor_get(x_15, 16); +x_46 = lean_ctor_get(x_15, 17); +x_47 = lean_ctor_get(x_15, 18); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_15); +x_48 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_48, 0, x_1); +lean_ctor_set(x_48, 1, x_2); +lean_ctor_set(x_48, 2, x_3); +lean_ctor_set_uint8(x_48, sizeof(void*)*3, x_4); +x_49 = lean_array_push(x_33, x_48); +x_50 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_50, 0, x_28); +lean_ctor_set(x_50, 1, x_29); +lean_ctor_set(x_50, 2, x_30); +lean_ctor_set(x_50, 3, x_31); +lean_ctor_set(x_50, 4, x_32); +lean_ctor_set(x_50, 5, x_49); +lean_ctor_set(x_50, 6, x_35); +lean_ctor_set(x_50, 7, x_36); +lean_ctor_set(x_50, 8, x_37); +lean_ctor_set(x_50, 9, x_38); +lean_ctor_set(x_50, 10, x_39); +lean_ctor_set(x_50, 11, x_40); +lean_ctor_set(x_50, 12, x_41); +lean_ctor_set(x_50, 13, x_42); +lean_ctor_set(x_50, 14, x_43); +lean_ctor_set(x_50, 15, x_44); +lean_ctor_set(x_50, 16, x_45); +lean_ctor_set(x_50, 17, x_46); +lean_ctor_set(x_50, 18, x_47); +lean_ctor_set_uint8(x_50, sizeof(void*)*19, x_34); +x_51 = lean_st_ref_set(x_5, x_50, x_16); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_53 = x_51; +} else { + lean_dec_ref(x_51); + x_53 = lean_box(0); +} +x_54 = lean_box(0); +if (lean_is_scalar(x_53)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_53; +} +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_52); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_4); +lean_dec(x_4); +x_15 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_15; +} +} +static uint64_t _init_l_Lean_Meta_Grind_hasSameType___closed__1() { +_start: +{ +uint8_t x_1; uint64_t x_2; +x_1 = 1; +x_2 = l_Lean_Meta_TransparencyMode_toUInt64(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_hasSameType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_3, 0); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +uint64_t x_11; uint8_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; lean_object* x_18; +x_11 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); +x_12 = 1; +lean_ctor_set_uint8(x_9, 9, x_12); +x_13 = 2; +x_14 = lean_uint64_shift_right(x_11, x_13); +x_15 = lean_uint64_shift_left(x_14, x_13); +x_16 = l_Lean_Meta_Grind_hasSameType___closed__1; +x_17 = lean_uint64_lor(x_15, x_16); +lean_ctor_set_uint64(x_3, sizeof(void*)*7, x_17); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_18 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_21 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Meta_isExprDefEq(x_19, x_22, x_3, x_4, x_5, x_6, x_23); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +return x_24; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_24); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +else +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) +{ +return x_24; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_24, 0); +x_31 = lean_ctor_get(x_24, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_24); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +else +{ +uint8_t x_33; +lean_dec(x_19); +lean_dec(x_3); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_33 = !lean_is_exclusive(x_21); +if (x_33 == 0) +{ +return x_21; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_21, 0); +x_35 = lean_ctor_get(x_21, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_21); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_3); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_18); +if (x_37 == 0) +{ +return x_18; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_18, 0); +x_39 = lean_ctor_get(x_18, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_18); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +else +{ +uint64_t x_41; uint8_t x_42; uint8_t x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; uint8_t x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; uint8_t x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; uint64_t x_60; uint64_t x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; lean_object* x_65; +x_41 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); +x_42 = lean_ctor_get_uint8(x_9, 0); +x_43 = lean_ctor_get_uint8(x_9, 1); +x_44 = lean_ctor_get_uint8(x_9, 2); +x_45 = lean_ctor_get_uint8(x_9, 3); +x_46 = lean_ctor_get_uint8(x_9, 4); +x_47 = lean_ctor_get_uint8(x_9, 5); +x_48 = lean_ctor_get_uint8(x_9, 6); +x_49 = lean_ctor_get_uint8(x_9, 7); +x_50 = lean_ctor_get_uint8(x_9, 8); +x_51 = lean_ctor_get_uint8(x_9, 10); +x_52 = lean_ctor_get_uint8(x_9, 11); +x_53 = lean_ctor_get_uint8(x_9, 12); +x_54 = lean_ctor_get_uint8(x_9, 13); +x_55 = lean_ctor_get_uint8(x_9, 14); +x_56 = lean_ctor_get_uint8(x_9, 15); +x_57 = lean_ctor_get_uint8(x_9, 16); +lean_dec(x_9); +x_58 = 1; +x_59 = lean_alloc_ctor(0, 0, 17); +lean_ctor_set_uint8(x_59, 0, x_42); +lean_ctor_set_uint8(x_59, 1, x_43); +lean_ctor_set_uint8(x_59, 2, x_44); +lean_ctor_set_uint8(x_59, 3, x_45); +lean_ctor_set_uint8(x_59, 4, x_46); +lean_ctor_set_uint8(x_59, 5, x_47); +lean_ctor_set_uint8(x_59, 6, x_48); +lean_ctor_set_uint8(x_59, 7, x_49); +lean_ctor_set_uint8(x_59, 8, x_50); +lean_ctor_set_uint8(x_59, 9, x_58); +lean_ctor_set_uint8(x_59, 10, x_51); +lean_ctor_set_uint8(x_59, 11, x_52); +lean_ctor_set_uint8(x_59, 12, x_53); +lean_ctor_set_uint8(x_59, 13, x_54); +lean_ctor_set_uint8(x_59, 14, x_55); +lean_ctor_set_uint8(x_59, 15, x_56); +lean_ctor_set_uint8(x_59, 16, x_57); +x_60 = 2; +x_61 = lean_uint64_shift_right(x_41, x_60); +x_62 = lean_uint64_shift_left(x_61, x_60); +x_63 = l_Lean_Meta_Grind_hasSameType___closed__1; +x_64 = lean_uint64_lor(x_62, x_63); +lean_ctor_set(x_3, 0, x_59); +lean_ctor_set_uint64(x_3, sizeof(void*)*7, x_64); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_65 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_65) == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_68 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_67); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_Meta_isExprDefEq(x_66, x_69, x_3, x_4, x_5, x_6, x_70); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_74 = x_71; +} else { + lean_dec_ref(x_71); + x_74 = lean_box(0); +} +if (lean_is_scalar(x_74)) { + x_75 = lean_alloc_ctor(0, 2, 0); +} else { + x_75 = x_74; +} +lean_ctor_set(x_75, 0, x_72); +lean_ctor_set(x_75, 1, x_73); +return x_75; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_71, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_71, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_78 = x_71; +} else { + lean_dec_ref(x_71); + x_78 = lean_box(0); +} +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); +} else { + x_79 = x_78; +} +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_66); +lean_dec(x_3); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_80 = lean_ctor_get(x_68, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_68, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_82 = x_68; +} else { + lean_dec_ref(x_68); + x_82 = lean_box(0); +} +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(1, 2, 0); +} else { + x_83 = x_82; +} +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_81); +return x_83; +} +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_3); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_84 = lean_ctor_get(x_65, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_65, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_86 = x_65; +} else { + lean_dec_ref(x_65); + x_86 = lean_box(0); } if (lean_is_scalar(x_86)) { x_87 = lean_alloc_ctor(1, 2, 0); } else { - x_87 = x_86; + x_87 = x_86; +} +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_85); +return x_87; +} +} +} +else +{ +lean_object* x_88; uint64_t x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; uint8_t x_98; uint8_t x_99; uint8_t x_100; uint8_t x_101; uint8_t x_102; uint8_t x_103; uint8_t x_104; uint8_t x_105; uint8_t x_106; uint8_t x_107; uint8_t x_108; uint8_t x_109; uint8_t x_110; uint8_t x_111; uint8_t x_112; uint8_t x_113; uint8_t x_114; lean_object* x_115; uint8_t x_116; lean_object* x_117; uint64_t x_118; uint64_t x_119; uint64_t x_120; uint64_t x_121; uint64_t x_122; lean_object* x_123; lean_object* x_124; +x_88 = lean_ctor_get(x_3, 0); +x_89 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); +x_90 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 8); +x_91 = lean_ctor_get(x_3, 1); +x_92 = lean_ctor_get(x_3, 2); +x_93 = lean_ctor_get(x_3, 3); +x_94 = lean_ctor_get(x_3, 4); +x_95 = lean_ctor_get(x_3, 5); +x_96 = lean_ctor_get(x_3, 6); +x_97 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 9); +x_98 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 10); +lean_inc(x_96); +lean_inc(x_95); +lean_inc(x_94); +lean_inc(x_93); +lean_inc(x_92); +lean_inc(x_91); +lean_inc(x_88); +lean_dec(x_3); +x_99 = lean_ctor_get_uint8(x_88, 0); +x_100 = lean_ctor_get_uint8(x_88, 1); +x_101 = lean_ctor_get_uint8(x_88, 2); +x_102 = lean_ctor_get_uint8(x_88, 3); +x_103 = lean_ctor_get_uint8(x_88, 4); +x_104 = lean_ctor_get_uint8(x_88, 5); +x_105 = lean_ctor_get_uint8(x_88, 6); +x_106 = lean_ctor_get_uint8(x_88, 7); +x_107 = lean_ctor_get_uint8(x_88, 8); +x_108 = lean_ctor_get_uint8(x_88, 10); +x_109 = lean_ctor_get_uint8(x_88, 11); +x_110 = lean_ctor_get_uint8(x_88, 12); +x_111 = lean_ctor_get_uint8(x_88, 13); +x_112 = lean_ctor_get_uint8(x_88, 14); +x_113 = lean_ctor_get_uint8(x_88, 15); +x_114 = lean_ctor_get_uint8(x_88, 16); +if (lean_is_exclusive(x_88)) { + x_115 = x_88; +} else { + lean_dec_ref(x_88); + x_115 = lean_box(0); +} +x_116 = 1; +if (lean_is_scalar(x_115)) { + x_117 = lean_alloc_ctor(0, 0, 17); +} else { + x_117 = x_115; +} +lean_ctor_set_uint8(x_117, 0, x_99); +lean_ctor_set_uint8(x_117, 1, x_100); +lean_ctor_set_uint8(x_117, 2, x_101); +lean_ctor_set_uint8(x_117, 3, x_102); +lean_ctor_set_uint8(x_117, 4, x_103); +lean_ctor_set_uint8(x_117, 5, x_104); +lean_ctor_set_uint8(x_117, 6, x_105); +lean_ctor_set_uint8(x_117, 7, x_106); +lean_ctor_set_uint8(x_117, 8, x_107); +lean_ctor_set_uint8(x_117, 9, x_116); +lean_ctor_set_uint8(x_117, 10, x_108); +lean_ctor_set_uint8(x_117, 11, x_109); +lean_ctor_set_uint8(x_117, 12, x_110); +lean_ctor_set_uint8(x_117, 13, x_111); +lean_ctor_set_uint8(x_117, 14, x_112); +lean_ctor_set_uint8(x_117, 15, x_113); +lean_ctor_set_uint8(x_117, 16, x_114); +x_118 = 2; +x_119 = lean_uint64_shift_right(x_89, x_118); +x_120 = lean_uint64_shift_left(x_119, x_118); +x_121 = l_Lean_Meta_Grind_hasSameType___closed__1; +x_122 = lean_uint64_lor(x_120, x_121); +x_123 = lean_alloc_ctor(0, 7, 11); +lean_ctor_set(x_123, 0, x_117); +lean_ctor_set(x_123, 1, x_91); +lean_ctor_set(x_123, 2, x_92); +lean_ctor_set(x_123, 3, x_93); +lean_ctor_set(x_123, 4, x_94); +lean_ctor_set(x_123, 5, x_95); +lean_ctor_set(x_123, 6, x_96); +lean_ctor_set_uint64(x_123, sizeof(void*)*7, x_122); +lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 8, x_90); +lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 9, x_97); +lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 10, x_98); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_123); +x_124 = lean_infer_type(x_1, x_123, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_123); +x_127 = lean_infer_type(x_2, x_123, x_4, x_5, x_6, x_126); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +lean_dec(x_127); +x_130 = l_Lean_Meta_isExprDefEq(x_125, x_128, x_123, x_4, x_5, x_6, x_129); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_133 = x_130; +} else { + lean_dec_ref(x_130); + x_133 = lean_box(0); +} +if (lean_is_scalar(x_133)) { + x_134 = lean_alloc_ctor(0, 2, 0); +} else { + x_134 = x_133; +} +lean_ctor_set(x_134, 0, x_131); +lean_ctor_set(x_134, 1, x_132); +return x_134; +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_135 = lean_ctor_get(x_130, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_130, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_137 = x_130; +} else { + lean_dec_ref(x_130); + x_137 = lean_box(0); +} +if (lean_is_scalar(x_137)) { + x_138 = lean_alloc_ctor(1, 2, 0); +} else { + x_138 = x_137; +} +lean_ctor_set(x_138, 0, x_135); +lean_ctor_set(x_138, 1, x_136); +return x_138; +} +} +else +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_125); +lean_dec(x_123); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_139 = lean_ctor_get(x_127, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_127, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_141 = x_127; +} else { + lean_dec_ref(x_127); + x_141 = lean_box(0); +} +if (lean_is_scalar(x_141)) { + x_142 = lean_alloc_ctor(1, 2, 0); +} else { + x_142 = x_141; +} +lean_ctor_set(x_142, 0, x_139); +lean_ctor_set(x_142, 1, x_140); +return x_142; +} +} +else +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_123); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_143 = lean_ctor_get(x_124, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_124, 1); +lean_inc(x_144); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_145 = x_124; +} else { + lean_dec_ref(x_124); + x_145 = lean_box(0); +} +if (lean_is_scalar(x_145)) { + x_146 = lean_alloc_ctor(1, 2, 0); +} else { + x_146 = x_145; +} +lean_ctor_set(x_146, 0, x_143); +lean_ctor_set(x_146, 1, x_144); +return x_146; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = 1; +x_18 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_18; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = 0; +x_21 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_13); +if (x_22 == 0) +{ +return x_13; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_13, 0); +x_24 = lean_ctor_get(x_13, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_13); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_pushEqHEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = 0; +x_14 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_pushEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = 1; +x_14 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_pushHEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_12 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = 0; +x_16 = l_Lean_Meta_Grind_pushEqCore(x_1, x_13, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_12 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = 0; +x_16 = l_Lean_Meta_Grind_pushEqCore(x_1, x_13, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_pushEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_box(0); +x_5 = 0; +x_6 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*4, x_5); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_1, 3); +x_13 = l_Lean_Expr_quickComp(x_2, x_10); +switch (x_13) { +case 0: +{ +lean_object* x_14; uint8_t x_15; +x_14 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_9, x_2, x_3); +x_15 = 0; +lean_ctor_set(x_1, 0, x_14); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_15); +return x_1; +} +case 1: +{ +uint8_t x_16; +lean_dec(x_11); +lean_dec(x_10); +x_16 = 0; +lean_ctor_set(x_1, 2, x_3); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); +return x_1; +} +default: +{ +lean_object* x_17; uint8_t x_18; +x_17 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_12, x_2, x_3); +x_18 = 0; +lean_ctor_set(x_1, 3, x_17); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_18); +return x_1; +} +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_1, 0); +x_20 = lean_ctor_get(x_1, 1); +x_21 = lean_ctor_get(x_1, 2); +x_22 = lean_ctor_get(x_1, 3); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_1); +x_23 = l_Lean_Expr_quickComp(x_2, x_20); +switch (x_23) { +case 0: +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_24 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_19, x_2, x_3); +x_25 = 0; +x_26 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_20); +lean_ctor_set(x_26, 2, x_21); +lean_ctor_set(x_26, 3, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*4, x_25); +return x_26; +} +case 1: +{ +uint8_t x_27; lean_object* x_28; +lean_dec(x_21); +lean_dec(x_20); +x_27 = 0; +x_28 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_28, 0, x_19); +lean_ctor_set(x_28, 1, x_2); +lean_ctor_set(x_28, 2, x_3); +lean_ctor_set(x_28, 3, x_22); +lean_ctor_set_uint8(x_28, sizeof(void*)*4, x_27); +return x_28; +} +default: +{ +lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_29 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_22, x_2, x_3); +x_30 = 0; +x_31 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_31, 0, x_19); +lean_ctor_set(x_31, 1, x_20); +lean_ctor_set(x_31, 2, x_21); +lean_ctor_set(x_31, 3, x_29); +lean_ctor_set_uint8(x_31, sizeof(void*)*4, x_30); +return x_31; +} +} +} +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_1); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_1, 0); +x_34 = lean_ctor_get(x_1, 1); +x_35 = lean_ctor_get(x_1, 2); +x_36 = lean_ctor_get(x_1, 3); +x_37 = l_Lean_Expr_quickComp(x_2, x_34); +switch (x_37) { +case 0: +{ +lean_object* x_38; uint8_t x_39; +x_38 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_33, x_2, x_3); +x_39 = lean_ctor_get_uint8(x_38, sizeof(void*)*4); +if (x_39 == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_38, 3); +lean_inc(x_41); +if (lean_obj_tag(x_41) == 0) +{ +uint8_t x_42; +x_42 = !lean_is_exclusive(x_38); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_ctor_get(x_38, 3); +lean_dec(x_43); +x_44 = lean_ctor_get(x_38, 0); +lean_dec(x_44); +lean_ctor_set(x_38, 0, x_41); +x_45 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_45); +return x_1; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = lean_ctor_get(x_38, 1); +x_47 = lean_ctor_get(x_38, 2); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_38); +x_48 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_48, 0, x_41); +lean_ctor_set(x_48, 1, x_46); +lean_ctor_set(x_48, 2, x_47); +lean_ctor_set(x_48, 3, x_41); +lean_ctor_set_uint8(x_48, sizeof(void*)*4, x_39); +x_49 = 1; +lean_ctor_set(x_1, 0, x_48); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_49); +return x_1; +} +} +else +{ +uint8_t x_50; +x_50 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); +if (x_50 == 0) +{ +uint8_t x_51; +x_51 = !lean_is_exclusive(x_38); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_38, 1); +x_53 = lean_ctor_get(x_38, 2); +x_54 = lean_ctor_get(x_38, 3); +lean_dec(x_54); +x_55 = lean_ctor_get(x_38, 0); +lean_dec(x_55); +x_56 = !lean_is_exclusive(x_41); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_62; +x_57 = lean_ctor_get(x_41, 0); +x_58 = lean_ctor_get(x_41, 1); +x_59 = lean_ctor_get(x_41, 2); +x_60 = lean_ctor_get(x_41, 3); +x_61 = 1; +lean_ctor_set(x_41, 3, x_57); +lean_ctor_set(x_41, 2, x_53); +lean_ctor_set(x_41, 1, x_52); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_61); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_60); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_61); +x_62 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_59); +lean_ctor_set(x_1, 1, x_58); +lean_ctor_set(x_1, 0, x_41); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_62); +return x_1; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; uint8_t x_69; +x_63 = lean_ctor_get(x_41, 0); +x_64 = lean_ctor_get(x_41, 1); +x_65 = lean_ctor_get(x_41, 2); +x_66 = lean_ctor_get(x_41, 3); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_41); +x_67 = 1; +x_68 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_52); +lean_ctor_set(x_68, 2, x_53); +lean_ctor_set(x_68, 3, x_63); +lean_ctor_set_uint8(x_68, sizeof(void*)*4, x_67); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_66); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_67); +x_69 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_65); +lean_ctor_set(x_1, 1, x_64); +lean_ctor_set(x_1, 0, x_68); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_69); +return x_1; +} +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_70 = lean_ctor_get(x_38, 1); +x_71 = lean_ctor_get(x_38, 2); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_38); +x_72 = lean_ctor_get(x_41, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_41, 1); +lean_inc(x_73); +x_74 = lean_ctor_get(x_41, 2); +lean_inc(x_74); +x_75 = lean_ctor_get(x_41, 3); +lean_inc(x_75); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + lean_ctor_release(x_41, 2); + lean_ctor_release(x_41, 3); + x_76 = x_41; +} else { + lean_dec_ref(x_41); + x_76 = lean_box(0); +} +x_77 = 1; +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(1, 4, 1); +} else { + x_78 = x_76; +} +lean_ctor_set(x_78, 0, x_40); +lean_ctor_set(x_78, 1, x_70); +lean_ctor_set(x_78, 2, x_71); +lean_ctor_set(x_78, 3, x_72); +lean_ctor_set_uint8(x_78, sizeof(void*)*4, x_77); +x_79 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_79, 0, x_75); +lean_ctor_set(x_79, 1, x_34); +lean_ctor_set(x_79, 2, x_35); +lean_ctor_set(x_79, 3, x_36); +lean_ctor_set_uint8(x_79, sizeof(void*)*4, x_77); +x_80 = 0; +lean_ctor_set(x_1, 3, x_79); +lean_ctor_set(x_1, 2, x_74); +lean_ctor_set(x_1, 1, x_73); +lean_ctor_set(x_1, 0, x_78); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_80); +return x_1; +} +} +else +{ +uint8_t x_81; +lean_free_object(x_1); +x_81 = !lean_is_exclusive(x_41); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_82 = lean_ctor_get(x_41, 3); +lean_dec(x_82); +x_83 = lean_ctor_get(x_41, 2); +lean_dec(x_83); +x_84 = lean_ctor_get(x_41, 1); +lean_dec(x_84); +x_85 = lean_ctor_get(x_41, 0); +lean_dec(x_85); +x_86 = 1; +lean_ctor_set(x_41, 3, x_36); +lean_ctor_set(x_41, 2, x_35); +lean_ctor_set(x_41, 1, x_34); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_86); +return x_41; +} +else +{ +uint8_t x_87; lean_object* x_88; +lean_dec(x_41); +x_87 = 1; +x_88 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_88, 0, x_38); +lean_ctor_set(x_88, 1, x_34); +lean_ctor_set(x_88, 2, x_35); +lean_ctor_set(x_88, 3, x_36); +lean_ctor_set_uint8(x_88, sizeof(void*)*4, x_87); +return x_88; +} +} +} +} +else +{ +uint8_t x_89; +x_89 = lean_ctor_get_uint8(x_40, sizeof(void*)*4); +if (x_89 == 0) +{ +uint8_t x_90; +x_90 = !lean_is_exclusive(x_38); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_91 = lean_ctor_get(x_38, 1); +x_92 = lean_ctor_get(x_38, 2); +x_93 = lean_ctor_get(x_38, 3); +x_94 = lean_ctor_get(x_38, 0); +lean_dec(x_94); +x_95 = !lean_is_exclusive(x_40); +if (x_95 == 0) +{ +uint8_t x_96; uint8_t x_97; +x_96 = 1; +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_96); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_93); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_96); +x_97 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_92); +lean_ctor_set(x_1, 1, x_91); +lean_ctor_set(x_1, 0, x_40); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_97); +return x_1; +} +else +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; uint8_t x_104; +x_98 = lean_ctor_get(x_40, 0); +x_99 = lean_ctor_get(x_40, 1); +x_100 = lean_ctor_get(x_40, 2); +x_101 = lean_ctor_get(x_40, 3); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_40); +x_102 = 1; +x_103 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_103, 0, x_98); +lean_ctor_set(x_103, 1, x_99); +lean_ctor_set(x_103, 2, x_100); +lean_ctor_set(x_103, 3, x_101); +lean_ctor_set_uint8(x_103, sizeof(void*)*4, x_102); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_93); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_102); +x_104 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_92); +lean_ctor_set(x_1, 1, x_91); +lean_ctor_set(x_1, 0, x_103); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_104); +return x_1; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_105 = lean_ctor_get(x_38, 1); +x_106 = lean_ctor_get(x_38, 2); +x_107 = lean_ctor_get(x_38, 3); +lean_inc(x_107); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_38); +x_108 = lean_ctor_get(x_40, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_40, 1); +lean_inc(x_109); +x_110 = lean_ctor_get(x_40, 2); +lean_inc(x_110); +x_111 = lean_ctor_get(x_40, 3); +lean_inc(x_111); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_112 = x_40; +} else { + lean_dec_ref(x_40); + x_112 = lean_box(0); +} +x_113 = 1; +if (lean_is_scalar(x_112)) { + x_114 = lean_alloc_ctor(1, 4, 1); +} else { + x_114 = x_112; +} +lean_ctor_set(x_114, 0, x_108); +lean_ctor_set(x_114, 1, x_109); +lean_ctor_set(x_114, 2, x_110); +lean_ctor_set(x_114, 3, x_111); +lean_ctor_set_uint8(x_114, sizeof(void*)*4, x_113); +x_115 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_115, 0, x_107); +lean_ctor_set(x_115, 1, x_34); +lean_ctor_set(x_115, 2, x_35); +lean_ctor_set(x_115, 3, x_36); +lean_ctor_set_uint8(x_115, sizeof(void*)*4, x_113); +x_116 = 0; +lean_ctor_set(x_1, 3, x_115); +lean_ctor_set(x_1, 2, x_106); +lean_ctor_set(x_1, 1, x_105); +lean_ctor_set(x_1, 0, x_114); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_116); +return x_1; +} +} +else +{ +lean_object* x_117; +x_117 = lean_ctor_get(x_38, 3); +lean_inc(x_117); +if (lean_obj_tag(x_117) == 0) +{ +uint8_t x_118; +lean_free_object(x_1); +x_118 = !lean_is_exclusive(x_40); +if (x_118 == 0) +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_119 = lean_ctor_get(x_40, 3); +lean_dec(x_119); +x_120 = lean_ctor_get(x_40, 2); +lean_dec(x_120); +x_121 = lean_ctor_get(x_40, 1); +lean_dec(x_121); +x_122 = lean_ctor_get(x_40, 0); +lean_dec(x_122); +x_123 = 1; +lean_ctor_set(x_40, 3, x_36); +lean_ctor_set(x_40, 2, x_35); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_123); +return x_40; +} +else +{ +uint8_t x_124; lean_object* x_125; +lean_dec(x_40); +x_124 = 1; +x_125 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_125, 0, x_38); +lean_ctor_set(x_125, 1, x_34); +lean_ctor_set(x_125, 2, x_35); +lean_ctor_set(x_125, 3, x_36); +lean_ctor_set_uint8(x_125, sizeof(void*)*4, x_124); +return x_125; +} +} +else +{ +uint8_t x_126; +x_126 = lean_ctor_get_uint8(x_117, sizeof(void*)*4); +if (x_126 == 0) +{ +uint8_t x_127; +lean_free_object(x_1); +x_127 = !lean_is_exclusive(x_38); +if (x_127 == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_128 = lean_ctor_get(x_38, 1); +x_129 = lean_ctor_get(x_38, 2); +x_130 = lean_ctor_get(x_38, 3); +lean_dec(x_130); +x_131 = lean_ctor_get(x_38, 0); +lean_dec(x_131); +x_132 = !lean_is_exclusive(x_117); +if (x_132 == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; +x_133 = lean_ctor_get(x_117, 0); +x_134 = lean_ctor_get(x_117, 1); +x_135 = lean_ctor_get(x_117, 2); +x_136 = lean_ctor_get(x_117, 3); +x_137 = 1; +lean_inc(x_40); +lean_ctor_set(x_117, 3, x_133); +lean_ctor_set(x_117, 2, x_129); +lean_ctor_set(x_117, 1, x_128); +lean_ctor_set(x_117, 0, x_40); +x_138 = !lean_is_exclusive(x_40); +if (x_138 == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; +x_139 = lean_ctor_get(x_40, 3); +lean_dec(x_139); +x_140 = lean_ctor_get(x_40, 2); +lean_dec(x_140); +x_141 = lean_ctor_get(x_40, 1); +lean_dec(x_141); +x_142 = lean_ctor_get(x_40, 0); +lean_dec(x_142); +lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); +lean_ctor_set(x_40, 3, x_36); +lean_ctor_set(x_40, 2, x_35); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 0, x_136); +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_137); +x_143 = 0; +lean_ctor_set(x_38, 3, x_40); +lean_ctor_set(x_38, 2, x_135); +lean_ctor_set(x_38, 1, x_134); +lean_ctor_set(x_38, 0, x_117); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_143); +return x_38; +} +else +{ +lean_object* x_144; uint8_t x_145; +lean_dec(x_40); +lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); +x_144 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_144, 0, x_136); +lean_ctor_set(x_144, 1, x_34); +lean_ctor_set(x_144, 2, x_35); +lean_ctor_set(x_144, 3, x_36); +lean_ctor_set_uint8(x_144, sizeof(void*)*4, x_137); +x_145 = 0; +lean_ctor_set(x_38, 3, x_144); +lean_ctor_set(x_38, 2, x_135); +lean_ctor_set(x_38, 1, x_134); +lean_ctor_set(x_38, 0, x_117); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_145); +return x_38; +} +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; +x_146 = lean_ctor_get(x_117, 0); +x_147 = lean_ctor_get(x_117, 1); +x_148 = lean_ctor_get(x_117, 2); +x_149 = lean_ctor_get(x_117, 3); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_117); +x_150 = 1; +lean_inc(x_40); +x_151 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_151, 0, x_40); +lean_ctor_set(x_151, 1, x_128); +lean_ctor_set(x_151, 2, x_129); +lean_ctor_set(x_151, 3, x_146); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_152 = x_40; +} else { + lean_dec_ref(x_40); + x_152 = lean_box(0); +} +lean_ctor_set_uint8(x_151, sizeof(void*)*4, x_150); +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 4, 1); +} else { + x_153 = x_152; +} +lean_ctor_set(x_153, 0, x_149); +lean_ctor_set(x_153, 1, x_34); +lean_ctor_set(x_153, 2, x_35); +lean_ctor_set(x_153, 3, x_36); +lean_ctor_set_uint8(x_153, sizeof(void*)*4, x_150); +x_154 = 0; +lean_ctor_set(x_38, 3, x_153); +lean_ctor_set(x_38, 2, x_148); +lean_ctor_set(x_38, 1, x_147); +lean_ctor_set(x_38, 0, x_151); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_154); +return x_38; +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; +x_155 = lean_ctor_get(x_38, 1); +x_156 = lean_ctor_get(x_38, 2); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_38); +x_157 = lean_ctor_get(x_117, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_117, 1); +lean_inc(x_158); +x_159 = lean_ctor_get(x_117, 2); +lean_inc(x_159); +x_160 = lean_ctor_get(x_117, 3); +lean_inc(x_160); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + lean_ctor_release(x_117, 2); + lean_ctor_release(x_117, 3); + x_161 = x_117; +} else { + lean_dec_ref(x_117); + x_161 = lean_box(0); +} +x_162 = 1; +lean_inc(x_40); +if (lean_is_scalar(x_161)) { + x_163 = lean_alloc_ctor(1, 4, 1); +} else { + x_163 = x_161; +} +lean_ctor_set(x_163, 0, x_40); +lean_ctor_set(x_163, 1, x_155); +lean_ctor_set(x_163, 2, x_156); +lean_ctor_set(x_163, 3, x_157); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_164 = x_40; +} else { + lean_dec_ref(x_40); + x_164 = lean_box(0); +} +lean_ctor_set_uint8(x_163, sizeof(void*)*4, x_162); +if (lean_is_scalar(x_164)) { + x_165 = lean_alloc_ctor(1, 4, 1); +} else { + x_165 = x_164; +} +lean_ctor_set(x_165, 0, x_160); +lean_ctor_set(x_165, 1, x_34); +lean_ctor_set(x_165, 2, x_35); +lean_ctor_set(x_165, 3, x_36); +lean_ctor_set_uint8(x_165, sizeof(void*)*4, x_162); +x_166 = 0; +x_167 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_167, 0, x_163); +lean_ctor_set(x_167, 1, x_158); +lean_ctor_set(x_167, 2, x_159); +lean_ctor_set(x_167, 3, x_165); +lean_ctor_set_uint8(x_167, sizeof(void*)*4, x_166); +return x_167; +} +} +else +{ +uint8_t x_168; +x_168 = !lean_is_exclusive(x_38); +if (x_168 == 0) +{ +lean_object* x_169; lean_object* x_170; uint8_t x_171; +x_169 = lean_ctor_get(x_38, 3); +lean_dec(x_169); +x_170 = lean_ctor_get(x_38, 0); +lean_dec(x_170); +x_171 = !lean_is_exclusive(x_40); +if (x_171 == 0) +{ +uint8_t x_172; +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_126); +x_172 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_172); +return x_1; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; +x_173 = lean_ctor_get(x_40, 0); +x_174 = lean_ctor_get(x_40, 1); +x_175 = lean_ctor_get(x_40, 2); +x_176 = lean_ctor_get(x_40, 3); +lean_inc(x_176); +lean_inc(x_175); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_40); +x_177 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_177, 0, x_173); +lean_ctor_set(x_177, 1, x_174); +lean_ctor_set(x_177, 2, x_175); +lean_ctor_set(x_177, 3, x_176); +lean_ctor_set_uint8(x_177, sizeof(void*)*4, x_126); +lean_ctor_set(x_38, 0, x_177); +x_178 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_178); +return x_1; +} +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; +x_179 = lean_ctor_get(x_38, 1); +x_180 = lean_ctor_get(x_38, 2); +lean_inc(x_180); +lean_inc(x_179); +lean_dec(x_38); +x_181 = lean_ctor_get(x_40, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_40, 1); +lean_inc(x_182); +x_183 = lean_ctor_get(x_40, 2); +lean_inc(x_183); +x_184 = lean_ctor_get(x_40, 3); +lean_inc(x_184); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_185 = x_40; +} else { + lean_dec_ref(x_40); + x_185 = lean_box(0); +} +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 4, 1); +} else { + x_186 = x_185; +} +lean_ctor_set(x_186, 0, x_181); +lean_ctor_set(x_186, 1, x_182); +lean_ctor_set(x_186, 2, x_183); +lean_ctor_set(x_186, 3, x_184); +lean_ctor_set_uint8(x_186, sizeof(void*)*4, x_126); +x_187 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_187, 0, x_186); +lean_ctor_set(x_187, 1, x_179); +lean_ctor_set(x_187, 2, x_180); +lean_ctor_set(x_187, 3, x_117); +lean_ctor_set_uint8(x_187, sizeof(void*)*4, x_39); +x_188 = 1; +lean_ctor_set(x_1, 0, x_187); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_188); +return x_1; +} +} +} +} +} +} +else +{ +uint8_t x_189; +x_189 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_189); +return x_1; +} +} +case 1: +{ +uint8_t x_190; +lean_dec(x_35); +lean_dec(x_34); +x_190 = 1; +lean_ctor_set(x_1, 2, x_3); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_190); +return x_1; +} +default: +{ +lean_object* x_191; uint8_t x_192; +x_191 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_36, x_2, x_3); +x_192 = lean_ctor_get_uint8(x_191, sizeof(void*)*4); +if (x_192 == 0) +{ +lean_object* x_193; +x_193 = lean_ctor_get(x_191, 0); +lean_inc(x_193); +if (lean_obj_tag(x_193) == 0) +{ +lean_object* x_194; +x_194 = lean_ctor_get(x_191, 3); +lean_inc(x_194); +if (lean_obj_tag(x_194) == 0) +{ +uint8_t x_195; +x_195 = !lean_is_exclusive(x_191); +if (x_195 == 0) +{ +lean_object* x_196; lean_object* x_197; uint8_t x_198; +x_196 = lean_ctor_get(x_191, 3); +lean_dec(x_196); +x_197 = lean_ctor_get(x_191, 0); +lean_dec(x_197); +lean_ctor_set(x_191, 0, x_194); +x_198 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_198); +return x_1; +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; uint8_t x_202; +x_199 = lean_ctor_get(x_191, 1); +x_200 = lean_ctor_get(x_191, 2); +lean_inc(x_200); +lean_inc(x_199); +lean_dec(x_191); +x_201 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_201, 0, x_194); +lean_ctor_set(x_201, 1, x_199); +lean_ctor_set(x_201, 2, x_200); +lean_ctor_set(x_201, 3, x_194); +lean_ctor_set_uint8(x_201, sizeof(void*)*4, x_192); +x_202 = 1; +lean_ctor_set(x_1, 3, x_201); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_202); +return x_1; +} +} +else +{ +uint8_t x_203; +x_203 = lean_ctor_get_uint8(x_194, sizeof(void*)*4); +if (x_203 == 0) +{ +uint8_t x_204; +x_204 = !lean_is_exclusive(x_191); +if (x_204 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; +x_205 = lean_ctor_get(x_191, 1); +x_206 = lean_ctor_get(x_191, 2); +x_207 = lean_ctor_get(x_191, 3); +lean_dec(x_207); +x_208 = lean_ctor_get(x_191, 0); +lean_dec(x_208); +x_209 = !lean_is_exclusive(x_194); +if (x_209 == 0) +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; uint8_t x_215; +x_210 = lean_ctor_get(x_194, 0); +x_211 = lean_ctor_get(x_194, 1); +x_212 = lean_ctor_get(x_194, 2); +x_213 = lean_ctor_get(x_194, 3); +x_214 = 1; +lean_ctor_set(x_194, 3, x_193); +lean_ctor_set(x_194, 2, x_35); +lean_ctor_set(x_194, 1, x_34); +lean_ctor_set(x_194, 0, x_33); +lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_214); +lean_ctor_set(x_191, 3, x_213); +lean_ctor_set(x_191, 2, x_212); +lean_ctor_set(x_191, 1, x_211); +lean_ctor_set(x_191, 0, x_210); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_214); +x_215 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_206); +lean_ctor_set(x_1, 1, x_205); +lean_ctor_set(x_1, 0, x_194); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_215); +return x_1; +} +else +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; uint8_t x_222; +x_216 = lean_ctor_get(x_194, 0); +x_217 = lean_ctor_get(x_194, 1); +x_218 = lean_ctor_get(x_194, 2); +x_219 = lean_ctor_get(x_194, 3); +lean_inc(x_219); +lean_inc(x_218); +lean_inc(x_217); +lean_inc(x_216); +lean_dec(x_194); +x_220 = 1; +x_221 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_221, 0, x_33); +lean_ctor_set(x_221, 1, x_34); +lean_ctor_set(x_221, 2, x_35); +lean_ctor_set(x_221, 3, x_193); +lean_ctor_set_uint8(x_221, sizeof(void*)*4, x_220); +lean_ctor_set(x_191, 3, x_219); +lean_ctor_set(x_191, 2, x_218); +lean_ctor_set(x_191, 1, x_217); +lean_ctor_set(x_191, 0, x_216); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_220); +x_222 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_206); +lean_ctor_set(x_1, 1, x_205); +lean_ctor_set(x_1, 0, x_221); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_222); +return x_1; +} +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; uint8_t x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; +x_223 = lean_ctor_get(x_191, 1); +x_224 = lean_ctor_get(x_191, 2); +lean_inc(x_224); +lean_inc(x_223); +lean_dec(x_191); +x_225 = lean_ctor_get(x_194, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_194, 1); +lean_inc(x_226); +x_227 = lean_ctor_get(x_194, 2); +lean_inc(x_227); +x_228 = lean_ctor_get(x_194, 3); +lean_inc(x_228); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + lean_ctor_release(x_194, 2); + lean_ctor_release(x_194, 3); + x_229 = x_194; +} else { + lean_dec_ref(x_194); + x_229 = lean_box(0); } -lean_ctor_set(x_87, 0, x_84); -lean_ctor_set(x_87, 1, x_85); -return x_87; +x_230 = 1; +if (lean_is_scalar(x_229)) { + x_231 = lean_alloc_ctor(1, 4, 1); +} else { + x_231 = x_229; } +lean_ctor_set(x_231, 0, x_33); +lean_ctor_set(x_231, 1, x_34); +lean_ctor_set(x_231, 2, x_35); +lean_ctor_set(x_231, 3, x_193); +lean_ctor_set_uint8(x_231, sizeof(void*)*4, x_230); +x_232 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_232, 0, x_225); +lean_ctor_set(x_232, 1, x_226); +lean_ctor_set(x_232, 2, x_227); +lean_ctor_set(x_232, 3, x_228); +lean_ctor_set_uint8(x_232, sizeof(void*)*4, x_230); +x_233 = 0; +lean_ctor_set(x_1, 3, x_232); +lean_ctor_set(x_1, 2, x_224); +lean_ctor_set(x_1, 1, x_223); +lean_ctor_set(x_1, 0, x_231); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_233); +return x_1; } } else { -lean_object* x_88; uint64_t x_89; uint8_t x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; uint8_t x_98; uint8_t x_99; uint8_t x_100; uint8_t x_101; uint8_t x_102; uint8_t x_103; uint8_t x_104; uint8_t x_105; uint8_t x_106; uint8_t x_107; uint8_t x_108; uint8_t x_109; uint8_t x_110; uint8_t x_111; uint8_t x_112; uint8_t x_113; uint8_t x_114; lean_object* x_115; uint8_t x_116; lean_object* x_117; uint64_t x_118; uint64_t x_119; uint64_t x_120; uint64_t x_121; uint64_t x_122; lean_object* x_123; lean_object* x_124; -x_88 = lean_ctor_get(x_3, 0); -x_89 = lean_ctor_get_uint64(x_3, sizeof(void*)*7); -x_90 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 8); -x_91 = lean_ctor_get(x_3, 1); -x_92 = lean_ctor_get(x_3, 2); -x_93 = lean_ctor_get(x_3, 3); -x_94 = lean_ctor_get(x_3, 4); -x_95 = lean_ctor_get(x_3, 5); -x_96 = lean_ctor_get(x_3, 6); -x_97 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 9); -x_98 = lean_ctor_get_uint8(x_3, sizeof(void*)*7 + 10); -lean_inc(x_96); -lean_inc(x_95); -lean_inc(x_94); -lean_inc(x_93); -lean_inc(x_92); -lean_inc(x_91); -lean_inc(x_88); -lean_dec(x_3); -x_99 = lean_ctor_get_uint8(x_88, 0); -x_100 = lean_ctor_get_uint8(x_88, 1); -x_101 = lean_ctor_get_uint8(x_88, 2); -x_102 = lean_ctor_get_uint8(x_88, 3); -x_103 = lean_ctor_get_uint8(x_88, 4); -x_104 = lean_ctor_get_uint8(x_88, 5); -x_105 = lean_ctor_get_uint8(x_88, 6); -x_106 = lean_ctor_get_uint8(x_88, 7); -x_107 = lean_ctor_get_uint8(x_88, 8); -x_108 = lean_ctor_get_uint8(x_88, 10); -x_109 = lean_ctor_get_uint8(x_88, 11); -x_110 = lean_ctor_get_uint8(x_88, 12); -x_111 = lean_ctor_get_uint8(x_88, 13); -x_112 = lean_ctor_get_uint8(x_88, 14); -x_113 = lean_ctor_get_uint8(x_88, 15); -x_114 = lean_ctor_get_uint8(x_88, 16); -if (lean_is_exclusive(x_88)) { - x_115 = x_88; -} else { - lean_dec_ref(x_88); - x_115 = lean_box(0); +uint8_t x_234; +lean_free_object(x_1); +x_234 = !lean_is_exclusive(x_194); +if (x_234 == 0) +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; +x_235 = lean_ctor_get(x_194, 3); +lean_dec(x_235); +x_236 = lean_ctor_get(x_194, 2); +lean_dec(x_236); +x_237 = lean_ctor_get(x_194, 1); +lean_dec(x_237); +x_238 = lean_ctor_get(x_194, 0); +lean_dec(x_238); +x_239 = 1; +lean_ctor_set(x_194, 3, x_191); +lean_ctor_set(x_194, 2, x_35); +lean_ctor_set(x_194, 1, x_34); +lean_ctor_set(x_194, 0, x_33); +lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_239); +return x_194; } -x_116 = 1; -if (lean_is_scalar(x_115)) { - x_117 = lean_alloc_ctor(0, 0, 17); -} else { - x_117 = x_115; +else +{ +uint8_t x_240; lean_object* x_241; +lean_dec(x_194); +x_240 = 1; +x_241 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_241, 0, x_33); +lean_ctor_set(x_241, 1, x_34); +lean_ctor_set(x_241, 2, x_35); +lean_ctor_set(x_241, 3, x_191); +lean_ctor_set_uint8(x_241, sizeof(void*)*4, x_240); +return x_241; } -lean_ctor_set_uint8(x_117, 0, x_99); -lean_ctor_set_uint8(x_117, 1, x_100); -lean_ctor_set_uint8(x_117, 2, x_101); -lean_ctor_set_uint8(x_117, 3, x_102); -lean_ctor_set_uint8(x_117, 4, x_103); -lean_ctor_set_uint8(x_117, 5, x_104); -lean_ctor_set_uint8(x_117, 6, x_105); -lean_ctor_set_uint8(x_117, 7, x_106); -lean_ctor_set_uint8(x_117, 8, x_107); -lean_ctor_set_uint8(x_117, 9, x_116); -lean_ctor_set_uint8(x_117, 10, x_108); -lean_ctor_set_uint8(x_117, 11, x_109); -lean_ctor_set_uint8(x_117, 12, x_110); -lean_ctor_set_uint8(x_117, 13, x_111); -lean_ctor_set_uint8(x_117, 14, x_112); -lean_ctor_set_uint8(x_117, 15, x_113); -lean_ctor_set_uint8(x_117, 16, x_114); -x_118 = 2; -x_119 = lean_uint64_shift_right(x_89, x_118); -x_120 = lean_uint64_shift_left(x_119, x_118); -x_121 = l_Lean_Meta_Grind_hasSameType___closed__1; -x_122 = lean_uint64_lor(x_120, x_121); -x_123 = lean_alloc_ctor(0, 7, 11); -lean_ctor_set(x_123, 0, x_117); -lean_ctor_set(x_123, 1, x_91); -lean_ctor_set(x_123, 2, x_92); -lean_ctor_set(x_123, 3, x_93); -lean_ctor_set(x_123, 4, x_94); -lean_ctor_set(x_123, 5, x_95); -lean_ctor_set(x_123, 6, x_96); -lean_ctor_set_uint64(x_123, sizeof(void*)*7, x_122); -lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 8, x_90); -lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 9, x_97); -lean_ctor_set_uint8(x_123, sizeof(void*)*7 + 10, x_98); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_123); -x_124 = lean_infer_type(x_1, x_123, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_124) == 0) +} +} +} +else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); -lean_inc(x_126); -lean_dec(x_124); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_123); -x_127 = lean_infer_type(x_2, x_123, x_4, x_5, x_6, x_126); -if (lean_obj_tag(x_127) == 0) +uint8_t x_242; +x_242 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); +if (x_242 == 0) { -lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_ctor_get(x_127, 1); -lean_inc(x_129); -lean_dec(x_127); -x_130 = l_Lean_Meta_isExprDefEq(x_125, x_128, x_123, x_4, x_5, x_6, x_129); -if (lean_obj_tag(x_130) == 0) +uint8_t x_243; +x_243 = !lean_is_exclusive(x_191); +if (x_243 == 0) { -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_130, 1); -lean_inc(x_132); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_133 = x_130; -} else { - lean_dec_ref(x_130); - x_133 = lean_box(0); +lean_object* x_244; uint8_t x_245; +x_244 = lean_ctor_get(x_191, 0); +lean_dec(x_244); +x_245 = !lean_is_exclusive(x_193); +if (x_245 == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; uint8_t x_251; +x_246 = lean_ctor_get(x_193, 0); +x_247 = lean_ctor_get(x_193, 1); +x_248 = lean_ctor_get(x_193, 2); +x_249 = lean_ctor_get(x_193, 3); +x_250 = 1; +lean_ctor_set(x_193, 3, x_246); +lean_ctor_set(x_193, 2, x_35); +lean_ctor_set(x_193, 1, x_34); +lean_ctor_set(x_193, 0, x_33); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_250); +lean_ctor_set(x_191, 0, x_249); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_250); +x_251 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_248); +lean_ctor_set(x_1, 1, x_247); +lean_ctor_set(x_1, 0, x_193); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_251); +return x_1; } -if (lean_is_scalar(x_133)) { - x_134 = lean_alloc_ctor(0, 2, 0); -} else { - x_134 = x_133; +else +{ +lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; uint8_t x_256; lean_object* x_257; uint8_t x_258; +x_252 = lean_ctor_get(x_193, 0); +x_253 = lean_ctor_get(x_193, 1); +x_254 = lean_ctor_get(x_193, 2); +x_255 = lean_ctor_get(x_193, 3); +lean_inc(x_255); +lean_inc(x_254); +lean_inc(x_253); +lean_inc(x_252); +lean_dec(x_193); +x_256 = 1; +x_257 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_257, 0, x_33); +lean_ctor_set(x_257, 1, x_34); +lean_ctor_set(x_257, 2, x_35); +lean_ctor_set(x_257, 3, x_252); +lean_ctor_set_uint8(x_257, sizeof(void*)*4, x_256); +lean_ctor_set(x_191, 0, x_255); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_256); +x_258 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_254); +lean_ctor_set(x_1, 1, x_253); +lean_ctor_set(x_1, 0, x_257); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_258); +return x_1; } -lean_ctor_set(x_134, 0, x_131); -lean_ctor_set(x_134, 1, x_132); -return x_134; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_135 = lean_ctor_get(x_130, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_130, 1); -lean_inc(x_136); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_137 = x_130; +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint8_t x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; +x_259 = lean_ctor_get(x_191, 1); +x_260 = lean_ctor_get(x_191, 2); +x_261 = lean_ctor_get(x_191, 3); +lean_inc(x_261); +lean_inc(x_260); +lean_inc(x_259); +lean_dec(x_191); +x_262 = lean_ctor_get(x_193, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_193, 1); +lean_inc(x_263); +x_264 = lean_ctor_get(x_193, 2); +lean_inc(x_264); +x_265 = lean_ctor_get(x_193, 3); +lean_inc(x_265); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_266 = x_193; } else { - lean_dec_ref(x_130); - x_137 = lean_box(0); + lean_dec_ref(x_193); + x_266 = lean_box(0); } -if (lean_is_scalar(x_137)) { - x_138 = lean_alloc_ctor(1, 2, 0); +x_267 = 1; +if (lean_is_scalar(x_266)) { + x_268 = lean_alloc_ctor(1, 4, 1); } else { - x_138 = x_137; + x_268 = x_266; } -lean_ctor_set(x_138, 0, x_135); -lean_ctor_set(x_138, 1, x_136); -return x_138; +lean_ctor_set(x_268, 0, x_33); +lean_ctor_set(x_268, 1, x_34); +lean_ctor_set(x_268, 2, x_35); +lean_ctor_set(x_268, 3, x_262); +lean_ctor_set_uint8(x_268, sizeof(void*)*4, x_267); +x_269 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_269, 0, x_265); +lean_ctor_set(x_269, 1, x_259); +lean_ctor_set(x_269, 2, x_260); +lean_ctor_set(x_269, 3, x_261); +lean_ctor_set_uint8(x_269, sizeof(void*)*4, x_267); +x_270 = 0; +lean_ctor_set(x_1, 3, x_269); +lean_ctor_set(x_1, 2, x_264); +lean_ctor_set(x_1, 1, x_263); +lean_ctor_set(x_1, 0, x_268); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_270); +return x_1; } } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_dec(x_125); -lean_dec(x_123); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_139 = lean_ctor_get(x_127, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_127, 1); -lean_inc(x_140); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - lean_ctor_release(x_127, 1); - x_141 = x_127; -} else { - lean_dec_ref(x_127); - x_141 = lean_box(0); -} -if (lean_is_scalar(x_141)) { - x_142 = lean_alloc_ctor(1, 2, 0); -} else { - x_142 = x_141; -} -lean_ctor_set(x_142, 0, x_139); -lean_ctor_set(x_142, 1, x_140); -return x_142; -} +lean_object* x_271; +x_271 = lean_ctor_get(x_191, 3); +lean_inc(x_271); +if (lean_obj_tag(x_271) == 0) +{ +uint8_t x_272; +lean_free_object(x_1); +x_272 = !lean_is_exclusive(x_193); +if (x_272 == 0) +{ +lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; uint8_t x_277; +x_273 = lean_ctor_get(x_193, 3); +lean_dec(x_273); +x_274 = lean_ctor_get(x_193, 2); +lean_dec(x_274); +x_275 = lean_ctor_get(x_193, 1); +lean_dec(x_275); +x_276 = lean_ctor_get(x_193, 0); +lean_dec(x_276); +x_277 = 1; +lean_ctor_set(x_193, 3, x_191); +lean_ctor_set(x_193, 2, x_35); +lean_ctor_set(x_193, 1, x_34); +lean_ctor_set(x_193, 0, x_33); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_277); +return x_193; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_123); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -x_143 = lean_ctor_get(x_124, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_124, 1); -lean_inc(x_144); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_145 = x_124; -} else { - lean_dec_ref(x_124); - x_145 = lean_box(0); -} -if (lean_is_scalar(x_145)) { - x_146 = lean_alloc_ctor(1, 2, 0); -} else { - x_146 = x_145; -} -lean_ctor_set(x_146, 0, x_143); -lean_ctor_set(x_146, 1, x_144); -return x_146; -} -} +uint8_t x_278; lean_object* x_279; +lean_dec(x_193); +x_278 = 1; +x_279 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_279, 0, x_33); +lean_ctor_set(x_279, 1, x_34); +lean_ctor_set(x_279, 2, x_35); +lean_ctor_set(x_279, 3, x_191); +lean_ctor_set_uint8(x_279, sizeof(void*)*4, x_278); +return x_279; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_2); -lean_inc(x_1); -x_13 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +uint8_t x_280; +x_280 = lean_ctor_get_uint8(x_271, sizeof(void*)*4); +if (x_280 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_unbox(x_14); -lean_dec(x_14); -if (x_15 == 0) +uint8_t x_281; +lean_free_object(x_1); +x_281 = !lean_is_exclusive(x_191); +if (x_281 == 0) { -lean_object* x_16; uint8_t x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = 1; -x_18 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_18; +lean_object* x_282; lean_object* x_283; uint8_t x_284; +x_282 = lean_ctor_get(x_191, 3); +lean_dec(x_282); +x_283 = lean_ctor_get(x_191, 0); +lean_dec(x_283); +x_284 = !lean_is_exclusive(x_271); +if (x_284 == 0) +{ +lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; uint8_t x_290; +x_285 = lean_ctor_get(x_271, 0); +x_286 = lean_ctor_get(x_271, 1); +x_287 = lean_ctor_get(x_271, 2); +x_288 = lean_ctor_get(x_271, 3); +x_289 = 1; +lean_inc(x_193); +lean_ctor_set(x_271, 3, x_193); +lean_ctor_set(x_271, 2, x_35); +lean_ctor_set(x_271, 1, x_34); +lean_ctor_set(x_271, 0, x_33); +x_290 = !lean_is_exclusive(x_193); +if (x_290 == 0) +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; +x_291 = lean_ctor_get(x_193, 3); +lean_dec(x_291); +x_292 = lean_ctor_get(x_193, 2); +lean_dec(x_292); +x_293 = lean_ctor_get(x_193, 1); +lean_dec(x_293); +x_294 = lean_ctor_get(x_193, 0); +lean_dec(x_294); +lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); +lean_ctor_set(x_193, 3, x_288); +lean_ctor_set(x_193, 2, x_287); +lean_ctor_set(x_193, 1, x_286); +lean_ctor_set(x_193, 0, x_285); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_289); +x_295 = 0; +lean_ctor_set(x_191, 3, x_193); +lean_ctor_set(x_191, 0, x_271); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_295); +return x_191; } else { -lean_object* x_19; uint8_t x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_dec(x_13); -x_20 = 0; -x_21 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_21; +lean_object* x_296; uint8_t x_297; +lean_dec(x_193); +lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); +x_296 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_296, 0, x_285); +lean_ctor_set(x_296, 1, x_286); +lean_ctor_set(x_296, 2, x_287); +lean_ctor_set(x_296, 3, x_288); +lean_ctor_set_uint8(x_296, sizeof(void*)*4, x_289); +x_297 = 0; +lean_ctor_set(x_191, 3, x_296); +lean_ctor_set(x_191, 0, x_271); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_297); +return x_191; } } else { -uint8_t x_22; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_13); -if (x_22 == 0) -{ -return x_13; +lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; uint8_t x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; +x_298 = lean_ctor_get(x_271, 0); +x_299 = lean_ctor_get(x_271, 1); +x_300 = lean_ctor_get(x_271, 2); +x_301 = lean_ctor_get(x_271, 3); +lean_inc(x_301); +lean_inc(x_300); +lean_inc(x_299); +lean_inc(x_298); +lean_dec(x_271); +x_302 = 1; +lean_inc(x_193); +x_303 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_303, 0, x_33); +lean_ctor_set(x_303, 1, x_34); +lean_ctor_set(x_303, 2, x_35); +lean_ctor_set(x_303, 3, x_193); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_304 = x_193; +} else { + lean_dec_ref(x_193); + x_304 = lean_box(0); +} +lean_ctor_set_uint8(x_303, sizeof(void*)*4, x_302); +if (lean_is_scalar(x_304)) { + x_305 = lean_alloc_ctor(1, 4, 1); +} else { + x_305 = x_304; +} +lean_ctor_set(x_305, 0, x_298); +lean_ctor_set(x_305, 1, x_299); +lean_ctor_set(x_305, 2, x_300); +lean_ctor_set(x_305, 3, x_301); +lean_ctor_set_uint8(x_305, sizeof(void*)*4, x_302); +x_306 = 0; +lean_ctor_set(x_191, 3, x_305); +lean_ctor_set(x_191, 0, x_303); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_306); +return x_191; +} } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_13, 0); -x_24 = lean_ctor_get(x_13, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_13); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; uint8_t x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; uint8_t x_318; lean_object* x_319; +x_307 = lean_ctor_get(x_191, 1); +x_308 = lean_ctor_get(x_191, 2); +lean_inc(x_308); +lean_inc(x_307); +lean_dec(x_191); +x_309 = lean_ctor_get(x_271, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_271, 1); +lean_inc(x_310); +x_311 = lean_ctor_get(x_271, 2); +lean_inc(x_311); +x_312 = lean_ctor_get(x_271, 3); +lean_inc(x_312); +if (lean_is_exclusive(x_271)) { + lean_ctor_release(x_271, 0); + lean_ctor_release(x_271, 1); + lean_ctor_release(x_271, 2); + lean_ctor_release(x_271, 3); + x_313 = x_271; +} else { + lean_dec_ref(x_271); + x_313 = lean_box(0); } +x_314 = 1; +lean_inc(x_193); +if (lean_is_scalar(x_313)) { + x_315 = lean_alloc_ctor(1, 4, 1); +} else { + x_315 = x_313; } +lean_ctor_set(x_315, 0, x_33); +lean_ctor_set(x_315, 1, x_34); +lean_ctor_set(x_315, 2, x_35); +lean_ctor_set(x_315, 3, x_193); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_316 = x_193; +} else { + lean_dec_ref(x_193); + x_316 = lean_box(0); } +lean_ctor_set_uint8(x_315, sizeof(void*)*4, x_314); +if (lean_is_scalar(x_316)) { + x_317 = lean_alloc_ctor(1, 4, 1); +} else { + x_317 = x_316; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqHEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_pushEqHEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; +lean_ctor_set(x_317, 0, x_309); +lean_ctor_set(x_317, 1, x_310); +lean_ctor_set(x_317, 2, x_311); +lean_ctor_set(x_317, 3, x_312); +lean_ctor_set_uint8(x_317, sizeof(void*)*4, x_314); +x_318 = 0; +x_319 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_319, 0, x_315); +lean_ctor_set(x_319, 1, x_307); +lean_ctor_set(x_319, 2, x_308); +lean_ctor_set(x_319, 3, x_317); +lean_ctor_set_uint8(x_319, sizeof(void*)*4, x_318); +return x_319; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -uint8_t x_13; lean_object* x_14; -x_13 = 0; -x_14 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} +uint8_t x_320; +x_320 = !lean_is_exclusive(x_191); +if (x_320 == 0) +{ +lean_object* x_321; lean_object* x_322; uint8_t x_323; +x_321 = lean_ctor_get(x_191, 3); +lean_dec(x_321); +x_322 = lean_ctor_get(x_191, 0); +lean_dec(x_322); +x_323 = !lean_is_exclusive(x_193); +if (x_323 == 0) +{ +uint8_t x_324; +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_280); +x_324 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_324); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_pushEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint8_t x_330; +x_325 = lean_ctor_get(x_193, 0); +x_326 = lean_ctor_get(x_193, 1); +x_327 = lean_ctor_get(x_193, 2); +x_328 = lean_ctor_get(x_193, 3); +lean_inc(x_328); +lean_inc(x_327); +lean_inc(x_326); +lean_inc(x_325); +lean_dec(x_193); +x_329 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_329, 0, x_325); +lean_ctor_set(x_329, 1, x_326); +lean_ctor_set(x_329, 2, x_327); +lean_ctor_set(x_329, 3, x_328); +lean_ctor_set_uint8(x_329, sizeof(void*)*4, x_280); +lean_ctor_set(x_191, 0, x_329); +x_330 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_330); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -uint8_t x_13; lean_object* x_14; -x_13 = 1; -x_14 = l_Lean_Meta_Grind_pushEqCore(x_1, x_2, x_3, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; uint8_t x_340; +x_331 = lean_ctor_get(x_191, 1); +x_332 = lean_ctor_get(x_191, 2); +lean_inc(x_332); +lean_inc(x_331); +lean_dec(x_191); +x_333 = lean_ctor_get(x_193, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_193, 1); +lean_inc(x_334); +x_335 = lean_ctor_get(x_193, 2); +lean_inc(x_335); +x_336 = lean_ctor_get(x_193, 3); +lean_inc(x_336); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_337 = x_193; +} else { + lean_dec_ref(x_193); + x_337 = lean_box(0); } +if (lean_is_scalar(x_337)) { + x_338 = lean_alloc_ctor(1, 4, 1); +} else { + x_338 = x_337; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushHEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_pushHEq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_13; +lean_ctor_set(x_338, 0, x_333); +lean_ctor_set(x_338, 1, x_334); +lean_ctor_set(x_338, 2, x_335); +lean_ctor_set(x_338, 3, x_336); +lean_ctor_set_uint8(x_338, sizeof(void*)*4, x_280); +x_339 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_339, 0, x_338); +lean_ctor_set(x_339, 1, x_331); +lean_ctor_set(x_339, 2, x_332); +lean_ctor_set(x_339, 3, x_271); +lean_ctor_set_uint8(x_339, sizeof(void*)*4, x_192); +x_340 = 1; +lean_ctor_set(x_1, 3, x_339); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_340); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_12 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = 0; -x_16 = l_Lean_Meta_Grind_pushEqCore(x_1, x_13, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqTrue___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_pushEqTrue(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_12 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = 0; -x_16 = l_Lean_Meta_Grind_pushEqCore(x_1, x_13, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_16; +uint8_t x_341; +x_341 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_341); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_pushEqFalse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_pushEqFalse(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -if (lean_obj_tag(x_1) == 0) +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; +x_342 = lean_ctor_get(x_1, 0); +x_343 = lean_ctor_get(x_1, 1); +x_344 = lean_ctor_get(x_1, 2); +x_345 = lean_ctor_get(x_1, 3); +lean_inc(x_345); +lean_inc(x_344); +lean_inc(x_343); +lean_inc(x_342); +lean_dec(x_1); +x_346 = l_Lean_Expr_quickComp(x_2, x_343); +switch (x_346) { +case 0: { -lean_object* x_4; uint8_t x_5; lean_object* x_6; -x_4 = lean_box(0); -x_5 = 0; -x_6 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_6, 0, x_4); -lean_ctor_set(x_6, 1, x_2); -lean_ctor_set(x_6, 2, x_3); -lean_ctor_set(x_6, 3, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*4, x_5); -return x_6; -} -else +lean_object* x_347; uint8_t x_348; +x_347 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_342, x_2, x_3); +x_348 = lean_ctor_get_uint8(x_347, sizeof(void*)*4); +if (x_348 == 0) { -uint8_t x_7; -x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); -if (x_7 == 0) +lean_object* x_349; +x_349 = lean_ctor_get(x_347, 0); +lean_inc(x_349); +if (lean_obj_tag(x_349) == 0) { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_1); -if (x_8 == 0) +lean_object* x_350; +x_350 = lean_ctor_get(x_347, 3); +lean_inc(x_350); +if (lean_obj_tag(x_350) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -x_11 = lean_ctor_get(x_1, 2); -x_12 = lean_ctor_get(x_1, 3); -x_13 = l_Lean_Expr_quickComp(x_2, x_10); -switch (x_13) { -case 0: +lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; uint8_t x_355; lean_object* x_356; +x_351 = lean_ctor_get(x_347, 1); +lean_inc(x_351); +x_352 = lean_ctor_get(x_347, 2); +lean_inc(x_352); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_353 = x_347; +} else { + lean_dec_ref(x_347); + x_353 = lean_box(0); +} +if (lean_is_scalar(x_353)) { + x_354 = lean_alloc_ctor(1, 4, 1); +} else { + x_354 = x_353; +} +lean_ctor_set(x_354, 0, x_350); +lean_ctor_set(x_354, 1, x_351); +lean_ctor_set(x_354, 2, x_352); +lean_ctor_set(x_354, 3, x_350); +lean_ctor_set_uint8(x_354, sizeof(void*)*4, x_348); +x_355 = 1; +x_356 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_356, 0, x_354); +lean_ctor_set(x_356, 1, x_343); +lean_ctor_set(x_356, 2, x_344); +lean_ctor_set(x_356, 3, x_345); +lean_ctor_set_uint8(x_356, sizeof(void*)*4, x_355); +return x_356; +} +else { -lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_9, x_2, x_3); -x_15 = 0; -lean_ctor_set(x_1, 0, x_14); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_15); -return x_1; +uint8_t x_357; +x_357 = lean_ctor_get_uint8(x_350, sizeof(void*)*4); +if (x_357 == 0) +{ +lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; uint8_t x_366; lean_object* x_367; lean_object* x_368; uint8_t x_369; lean_object* x_370; +x_358 = lean_ctor_get(x_347, 1); +lean_inc(x_358); +x_359 = lean_ctor_get(x_347, 2); +lean_inc(x_359); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_360 = x_347; +} else { + lean_dec_ref(x_347); + x_360 = lean_box(0); } -case 1: +x_361 = lean_ctor_get(x_350, 0); +lean_inc(x_361); +x_362 = lean_ctor_get(x_350, 1); +lean_inc(x_362); +x_363 = lean_ctor_get(x_350, 2); +lean_inc(x_363); +x_364 = lean_ctor_get(x_350, 3); +lean_inc(x_364); +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + lean_ctor_release(x_350, 2); + lean_ctor_release(x_350, 3); + x_365 = x_350; +} else { + lean_dec_ref(x_350); + x_365 = lean_box(0); +} +x_366 = 1; +if (lean_is_scalar(x_365)) { + x_367 = lean_alloc_ctor(1, 4, 1); +} else { + x_367 = x_365; +} +lean_ctor_set(x_367, 0, x_349); +lean_ctor_set(x_367, 1, x_358); +lean_ctor_set(x_367, 2, x_359); +lean_ctor_set(x_367, 3, x_361); +lean_ctor_set_uint8(x_367, sizeof(void*)*4, x_366); +if (lean_is_scalar(x_360)) { + x_368 = lean_alloc_ctor(1, 4, 1); +} else { + x_368 = x_360; +} +lean_ctor_set(x_368, 0, x_364); +lean_ctor_set(x_368, 1, x_343); +lean_ctor_set(x_368, 2, x_344); +lean_ctor_set(x_368, 3, x_345); +lean_ctor_set_uint8(x_368, sizeof(void*)*4, x_366); +x_369 = 0; +x_370 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_370, 0, x_367); +lean_ctor_set(x_370, 1, x_362); +lean_ctor_set(x_370, 2, x_363); +lean_ctor_set(x_370, 3, x_368); +lean_ctor_set_uint8(x_370, sizeof(void*)*4, x_369); +return x_370; +} +else { -uint8_t x_16; -lean_dec(x_11); -lean_dec(x_10); -x_16 = 0; -lean_ctor_set(x_1, 2, x_3); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); -return x_1; +lean_object* x_371; uint8_t x_372; lean_object* x_373; +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + lean_ctor_release(x_350, 2); + lean_ctor_release(x_350, 3); + x_371 = x_350; +} else { + lean_dec_ref(x_350); + x_371 = lean_box(0); +} +x_372 = 1; +if (lean_is_scalar(x_371)) { + x_373 = lean_alloc_ctor(1, 4, 1); +} else { + x_373 = x_371; } -default: -{ -lean_object* x_17; uint8_t x_18; -x_17 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_12, x_2, x_3); -x_18 = 0; -lean_ctor_set(x_1, 3, x_17); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_18); -return x_1; +lean_ctor_set(x_373, 0, x_347); +lean_ctor_set(x_373, 1, x_343); +lean_ctor_set(x_373, 2, x_344); +lean_ctor_set(x_373, 3, x_345); +lean_ctor_set_uint8(x_373, sizeof(void*)*4, x_372); +return x_373; } } } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_1, 0); -x_20 = lean_ctor_get(x_1, 1); -x_21 = lean_ctor_get(x_1, 2); -x_22 = lean_ctor_get(x_1, 3); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_1); -x_23 = l_Lean_Expr_quickComp(x_2, x_20); -switch (x_23) { -case 0: -{ -lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_24 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_19, x_2, x_3); -x_25 = 0; -x_26 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_20); -lean_ctor_set(x_26, 2, x_21); -lean_ctor_set(x_26, 3, x_22); -lean_ctor_set_uint8(x_26, sizeof(void*)*4, x_25); -return x_26; -} -case 1: +uint8_t x_374; +x_374 = lean_ctor_get_uint8(x_349, sizeof(void*)*4); +if (x_374 == 0) { -uint8_t x_27; lean_object* x_28; -lean_dec(x_21); -lean_dec(x_20); -x_27 = 0; -x_28 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_28, 0, x_19); -lean_ctor_set(x_28, 1, x_2); -lean_ctor_set(x_28, 2, x_3); -lean_ctor_set(x_28, 3, x_22); -lean_ctor_set_uint8(x_28, sizeof(void*)*4, x_27); -return x_28; +lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; uint8_t x_384; lean_object* x_385; lean_object* x_386; uint8_t x_387; lean_object* x_388; +x_375 = lean_ctor_get(x_347, 1); +lean_inc(x_375); +x_376 = lean_ctor_get(x_347, 2); +lean_inc(x_376); +x_377 = lean_ctor_get(x_347, 3); +lean_inc(x_377); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_378 = x_347; +} else { + lean_dec_ref(x_347); + x_378 = lean_box(0); } -default: -{ -lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_29 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_22, x_2, x_3); -x_30 = 0; -x_31 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_31, 0, x_19); -lean_ctor_set(x_31, 1, x_20); -lean_ctor_set(x_31, 2, x_21); -lean_ctor_set(x_31, 3, x_29); -lean_ctor_set_uint8(x_31, sizeof(void*)*4, x_30); -return x_31; +x_379 = lean_ctor_get(x_349, 0); +lean_inc(x_379); +x_380 = lean_ctor_get(x_349, 1); +lean_inc(x_380); +x_381 = lean_ctor_get(x_349, 2); +lean_inc(x_381); +x_382 = lean_ctor_get(x_349, 3); +lean_inc(x_382); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_383 = x_349; +} else { + lean_dec_ref(x_349); + x_383 = lean_box(0); } +x_384 = 1; +if (lean_is_scalar(x_383)) { + x_385 = lean_alloc_ctor(1, 4, 1); +} else { + x_385 = x_383; } +lean_ctor_set(x_385, 0, x_379); +lean_ctor_set(x_385, 1, x_380); +lean_ctor_set(x_385, 2, x_381); +lean_ctor_set(x_385, 3, x_382); +lean_ctor_set_uint8(x_385, sizeof(void*)*4, x_384); +if (lean_is_scalar(x_378)) { + x_386 = lean_alloc_ctor(1, 4, 1); +} else { + x_386 = x_378; } +lean_ctor_set(x_386, 0, x_377); +lean_ctor_set(x_386, 1, x_343); +lean_ctor_set(x_386, 2, x_344); +lean_ctor_set(x_386, 3, x_345); +lean_ctor_set_uint8(x_386, sizeof(void*)*4, x_384); +x_387 = 0; +x_388 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_388, 0, x_385); +lean_ctor_set(x_388, 1, x_375); +lean_ctor_set(x_388, 2, x_376); +lean_ctor_set(x_388, 3, x_386); +lean_ctor_set_uint8(x_388, sizeof(void*)*4, x_387); +return x_388; } else { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_1); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_1, 0); -x_34 = lean_ctor_get(x_1, 1); -x_35 = lean_ctor_get(x_1, 2); -x_36 = lean_ctor_get(x_1, 3); -x_37 = l_Lean_Expr_quickComp(x_2, x_34); -switch (x_37) { -case 0: -{ -lean_object* x_38; uint8_t x_39; -x_38 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_33, x_2, x_3); -x_39 = lean_ctor_get_uint8(x_38, sizeof(void*)*4); -if (x_39 == 0) -{ -lean_object* x_40; -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; -x_41 = lean_ctor_get(x_38, 3); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) -{ -uint8_t x_42; -x_42 = !lean_is_exclusive(x_38); -if (x_42 == 0) +lean_object* x_389; +x_389 = lean_ctor_get(x_347, 3); +lean_inc(x_389); +if (lean_obj_tag(x_389) == 0) { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_38, 3); -lean_dec(x_43); -x_44 = lean_ctor_get(x_38, 0); -lean_dec(x_44); -lean_ctor_set(x_38, 0, x_41); -x_45 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_45); -return x_1; +lean_object* x_390; uint8_t x_391; lean_object* x_392; +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_390 = x_349; +} else { + lean_dec_ref(x_349); + x_390 = lean_box(0); +} +x_391 = 1; +if (lean_is_scalar(x_390)) { + x_392 = lean_alloc_ctor(1, 4, 1); +} else { + x_392 = x_390; +} +lean_ctor_set(x_392, 0, x_347); +lean_ctor_set(x_392, 1, x_343); +lean_ctor_set(x_392, 2, x_344); +lean_ctor_set(x_392, 3, x_345); +lean_ctor_set_uint8(x_392, sizeof(void*)*4, x_391); +return x_392; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_46 = lean_ctor_get(x_38, 1); -x_47 = lean_ctor_get(x_38, 2); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_38); -x_48 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_48, 0, x_41); -lean_ctor_set(x_48, 1, x_46); -lean_ctor_set(x_48, 2, x_47); -lean_ctor_set(x_48, 3, x_41); -lean_ctor_set_uint8(x_48, sizeof(void*)*4, x_39); -x_49 = 1; -lean_ctor_set(x_1, 0, x_48); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_49); -return x_1; +uint8_t x_393; +x_393 = lean_ctor_get_uint8(x_389, sizeof(void*)*4); +if (x_393 == 0) +{ +lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; uint8_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; uint8_t x_406; lean_object* x_407; +x_394 = lean_ctor_get(x_347, 1); +lean_inc(x_394); +x_395 = lean_ctor_get(x_347, 2); +lean_inc(x_395); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_396 = x_347; +} else { + lean_dec_ref(x_347); + x_396 = lean_box(0); +} +x_397 = lean_ctor_get(x_389, 0); +lean_inc(x_397); +x_398 = lean_ctor_get(x_389, 1); +lean_inc(x_398); +x_399 = lean_ctor_get(x_389, 2); +lean_inc(x_399); +x_400 = lean_ctor_get(x_389, 3); +lean_inc(x_400); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + lean_ctor_release(x_389, 2); + lean_ctor_release(x_389, 3); + x_401 = x_389; +} else { + lean_dec_ref(x_389); + x_401 = lean_box(0); +} +x_402 = 1; +lean_inc(x_349); +if (lean_is_scalar(x_401)) { + x_403 = lean_alloc_ctor(1, 4, 1); +} else { + x_403 = x_401; +} +lean_ctor_set(x_403, 0, x_349); +lean_ctor_set(x_403, 1, x_394); +lean_ctor_set(x_403, 2, x_395); +lean_ctor_set(x_403, 3, x_397); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_404 = x_349; +} else { + lean_dec_ref(x_349); + x_404 = lean_box(0); } +lean_ctor_set_uint8(x_403, sizeof(void*)*4, x_402); +if (lean_is_scalar(x_404)) { + x_405 = lean_alloc_ctor(1, 4, 1); +} else { + x_405 = x_404; } -else -{ -uint8_t x_50; -x_50 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); -if (x_50 == 0) -{ -uint8_t x_51; -x_51 = !lean_is_exclusive(x_38); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_52 = lean_ctor_get(x_38, 1); -x_53 = lean_ctor_get(x_38, 2); -x_54 = lean_ctor_get(x_38, 3); -lean_dec(x_54); -x_55 = lean_ctor_get(x_38, 0); -lean_dec(x_55); -x_56 = !lean_is_exclusive(x_41); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_62; -x_57 = lean_ctor_get(x_41, 0); -x_58 = lean_ctor_get(x_41, 1); -x_59 = lean_ctor_get(x_41, 2); -x_60 = lean_ctor_get(x_41, 3); -x_61 = 1; -lean_ctor_set(x_41, 3, x_57); -lean_ctor_set(x_41, 2, x_53); -lean_ctor_set(x_41, 1, x_52); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_61); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_60); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_61); -x_62 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_59); -lean_ctor_set(x_1, 1, x_58); -lean_ctor_set(x_1, 0, x_41); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_62); -return x_1; +lean_ctor_set(x_405, 0, x_400); +lean_ctor_set(x_405, 1, x_343); +lean_ctor_set(x_405, 2, x_344); +lean_ctor_set(x_405, 3, x_345); +lean_ctor_set_uint8(x_405, sizeof(void*)*4, x_402); +x_406 = 0; +if (lean_is_scalar(x_396)) { + x_407 = lean_alloc_ctor(1, 4, 1); +} else { + x_407 = x_396; +} +lean_ctor_set(x_407, 0, x_403); +lean_ctor_set(x_407, 1, x_398); +lean_ctor_set(x_407, 2, x_399); +lean_ctor_set(x_407, 3, x_405); +lean_ctor_set_uint8(x_407, sizeof(void*)*4, x_406); +return x_407; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; uint8_t x_69; -x_63 = lean_ctor_get(x_41, 0); -x_64 = lean_ctor_get(x_41, 1); -x_65 = lean_ctor_get(x_41, 2); -x_66 = lean_ctor_get(x_41, 3); -lean_inc(x_66); -lean_inc(x_65); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_41); -x_67 = 1; -x_68 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_68, 0, x_40); -lean_ctor_set(x_68, 1, x_52); -lean_ctor_set(x_68, 2, x_53); -lean_ctor_set(x_68, 3, x_63); -lean_ctor_set_uint8(x_68, sizeof(void*)*4, x_67); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_66); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_67); -x_69 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_65); -lean_ctor_set(x_1, 1, x_64); -lean_ctor_set(x_1, 0, x_68); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_69); -return x_1; +lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; uint8_t x_418; lean_object* x_419; +x_408 = lean_ctor_get(x_347, 1); +lean_inc(x_408); +x_409 = lean_ctor_get(x_347, 2); +lean_inc(x_409); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_410 = x_347; +} else { + lean_dec_ref(x_347); + x_410 = lean_box(0); } +x_411 = lean_ctor_get(x_349, 0); +lean_inc(x_411); +x_412 = lean_ctor_get(x_349, 1); +lean_inc(x_412); +x_413 = lean_ctor_get(x_349, 2); +lean_inc(x_413); +x_414 = lean_ctor_get(x_349, 3); +lean_inc(x_414); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_415 = x_349; +} else { + lean_dec_ref(x_349); + x_415 = lean_box(0); } -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_70 = lean_ctor_get(x_38, 1); -x_71 = lean_ctor_get(x_38, 2); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_38); -x_72 = lean_ctor_get(x_41, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_41, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_41, 2); -lean_inc(x_74); -x_75 = lean_ctor_get(x_41, 3); -lean_inc(x_75); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - lean_ctor_release(x_41, 2); - lean_ctor_release(x_41, 3); - x_76 = x_41; +if (lean_is_scalar(x_415)) { + x_416 = lean_alloc_ctor(1, 4, 1); } else { - lean_dec_ref(x_41); - x_76 = lean_box(0); + x_416 = x_415; } -x_77 = 1; -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_416, 0, x_411); +lean_ctor_set(x_416, 1, x_412); +lean_ctor_set(x_416, 2, x_413); +lean_ctor_set(x_416, 3, x_414); +lean_ctor_set_uint8(x_416, sizeof(void*)*4, x_393); +if (lean_is_scalar(x_410)) { + x_417 = lean_alloc_ctor(1, 4, 1); } else { - x_78 = x_76; + x_417 = x_410; +} +lean_ctor_set(x_417, 0, x_416); +lean_ctor_set(x_417, 1, x_408); +lean_ctor_set(x_417, 2, x_409); +lean_ctor_set(x_417, 3, x_389); +lean_ctor_set_uint8(x_417, sizeof(void*)*4, x_348); +x_418 = 1; +x_419 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_419, 0, x_417); +lean_ctor_set(x_419, 1, x_343); +lean_ctor_set(x_419, 2, x_344); +lean_ctor_set(x_419, 3, x_345); +lean_ctor_set_uint8(x_419, sizeof(void*)*4, x_418); +return x_419; +} } -lean_ctor_set(x_78, 0, x_40); -lean_ctor_set(x_78, 1, x_70); -lean_ctor_set(x_78, 2, x_71); -lean_ctor_set(x_78, 3, x_72); -lean_ctor_set_uint8(x_78, sizeof(void*)*4, x_77); -x_79 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_79, 0, x_75); -lean_ctor_set(x_79, 1, x_34); -lean_ctor_set(x_79, 2, x_35); -lean_ctor_set(x_79, 3, x_36); -lean_ctor_set_uint8(x_79, sizeof(void*)*4, x_77); -x_80 = 0; -lean_ctor_set(x_1, 3, x_79); -lean_ctor_set(x_1, 2, x_74); -lean_ctor_set(x_1, 1, x_73); -lean_ctor_set(x_1, 0, x_78); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_80); -return x_1; } } -else -{ -uint8_t x_81; -lean_free_object(x_1); -x_81 = !lean_is_exclusive(x_41); -if (x_81 == 0) -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_82 = lean_ctor_get(x_41, 3); -lean_dec(x_82); -x_83 = lean_ctor_get(x_41, 2); -lean_dec(x_83); -x_84 = lean_ctor_get(x_41, 1); -lean_dec(x_84); -x_85 = lean_ctor_get(x_41, 0); -lean_dec(x_85); -x_86 = 1; -lean_ctor_set(x_41, 3, x_36); -lean_ctor_set(x_41, 2, x_35); -lean_ctor_set(x_41, 1, x_34); -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_86); -return x_41; } else { -uint8_t x_87; lean_object* x_88; -lean_dec(x_41); -x_87 = 1; -x_88 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_88, 0, x_38); -lean_ctor_set(x_88, 1, x_34); -lean_ctor_set(x_88, 2, x_35); -lean_ctor_set(x_88, 3, x_36); -lean_ctor_set_uint8(x_88, sizeof(void*)*4, x_87); -return x_88; -} +uint8_t x_420; lean_object* x_421; +x_420 = 1; +x_421 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_421, 0, x_347); +lean_ctor_set(x_421, 1, x_343); +lean_ctor_set(x_421, 2, x_344); +lean_ctor_set(x_421, 3, x_345); +lean_ctor_set_uint8(x_421, sizeof(void*)*4, x_420); +return x_421; } } +case 1: +{ +uint8_t x_422; lean_object* x_423; +lean_dec(x_344); +lean_dec(x_343); +x_422 = 1; +x_423 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_423, 0, x_342); +lean_ctor_set(x_423, 1, x_2); +lean_ctor_set(x_423, 2, x_3); +lean_ctor_set(x_423, 3, x_345); +lean_ctor_set_uint8(x_423, sizeof(void*)*4, x_422); +return x_423; } -else +default: { -uint8_t x_89; -x_89 = lean_ctor_get_uint8(x_40, sizeof(void*)*4); -if (x_89 == 0) +lean_object* x_424; uint8_t x_425; +x_424 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_345, x_2, x_3); +x_425 = lean_ctor_get_uint8(x_424, sizeof(void*)*4); +if (x_425 == 0) { -uint8_t x_90; -x_90 = !lean_is_exclusive(x_38); -if (x_90 == 0) +lean_object* x_426; +x_426 = lean_ctor_get(x_424, 0); +lean_inc(x_426); +if (lean_obj_tag(x_426) == 0) { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_91 = lean_ctor_get(x_38, 1); -x_92 = lean_ctor_get(x_38, 2); -x_93 = lean_ctor_get(x_38, 3); -x_94 = lean_ctor_get(x_38, 0); -lean_dec(x_94); -x_95 = !lean_is_exclusive(x_40); -if (x_95 == 0) +lean_object* x_427; +x_427 = lean_ctor_get(x_424, 3); +lean_inc(x_427); +if (lean_obj_tag(x_427) == 0) { -uint8_t x_96; uint8_t x_97; -x_96 = 1; -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_96); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_93); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_96); -x_97 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_92); -lean_ctor_set(x_1, 1, x_91); -lean_ctor_set(x_1, 0, x_40); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_97); -return x_1; +lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; uint8_t x_432; lean_object* x_433; +x_428 = lean_ctor_get(x_424, 1); +lean_inc(x_428); +x_429 = lean_ctor_get(x_424, 2); +lean_inc(x_429); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_430 = x_424; +} else { + lean_dec_ref(x_424); + x_430 = lean_box(0); } -else -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; uint8_t x_104; -x_98 = lean_ctor_get(x_40, 0); -x_99 = lean_ctor_get(x_40, 1); -x_100 = lean_ctor_get(x_40, 2); -x_101 = lean_ctor_get(x_40, 3); -lean_inc(x_101); -lean_inc(x_100); -lean_inc(x_99); -lean_inc(x_98); -lean_dec(x_40); -x_102 = 1; -x_103 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_103, 0, x_98); -lean_ctor_set(x_103, 1, x_99); -lean_ctor_set(x_103, 2, x_100); -lean_ctor_set(x_103, 3, x_101); -lean_ctor_set_uint8(x_103, sizeof(void*)*4, x_102); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_93); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_102); -x_104 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_92); -lean_ctor_set(x_1, 1, x_91); -lean_ctor_set(x_1, 0, x_103); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_104); -return x_1; +if (lean_is_scalar(x_430)) { + x_431 = lean_alloc_ctor(1, 4, 1); +} else { + x_431 = x_430; } +lean_ctor_set(x_431, 0, x_427); +lean_ctor_set(x_431, 1, x_428); +lean_ctor_set(x_431, 2, x_429); +lean_ctor_set(x_431, 3, x_427); +lean_ctor_set_uint8(x_431, sizeof(void*)*4, x_425); +x_432 = 1; +x_433 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_433, 0, x_342); +lean_ctor_set(x_433, 1, x_343); +lean_ctor_set(x_433, 2, x_344); +lean_ctor_set(x_433, 3, x_431); +lean_ctor_set_uint8(x_433, sizeof(void*)*4, x_432); +return x_433; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; -x_105 = lean_ctor_get(x_38, 1); -x_106 = lean_ctor_get(x_38, 2); -x_107 = lean_ctor_get(x_38, 3); -lean_inc(x_107); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_38); -x_108 = lean_ctor_get(x_40, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_40, 1); -lean_inc(x_109); -x_110 = lean_ctor_get(x_40, 2); -lean_inc(x_110); -x_111 = lean_ctor_get(x_40, 3); -lean_inc(x_111); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_112 = x_40; +uint8_t x_434; +x_434 = lean_ctor_get_uint8(x_427, sizeof(void*)*4); +if (x_434 == 0) +{ +lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; uint8_t x_443; lean_object* x_444; lean_object* x_445; uint8_t x_446; lean_object* x_447; +x_435 = lean_ctor_get(x_424, 1); +lean_inc(x_435); +x_436 = lean_ctor_get(x_424, 2); +lean_inc(x_436); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_437 = x_424; } else { - lean_dec_ref(x_40); - x_112 = lean_box(0); + lean_dec_ref(x_424); + x_437 = lean_box(0); } -x_113 = 1; -if (lean_is_scalar(x_112)) { - x_114 = lean_alloc_ctor(1, 4, 1); +x_438 = lean_ctor_get(x_427, 0); +lean_inc(x_438); +x_439 = lean_ctor_get(x_427, 1); +lean_inc(x_439); +x_440 = lean_ctor_get(x_427, 2); +lean_inc(x_440); +x_441 = lean_ctor_get(x_427, 3); +lean_inc(x_441); +if (lean_is_exclusive(x_427)) { + lean_ctor_release(x_427, 0); + lean_ctor_release(x_427, 1); + lean_ctor_release(x_427, 2); + lean_ctor_release(x_427, 3); + x_442 = x_427; } else { - x_114 = x_112; + lean_dec_ref(x_427); + x_442 = lean_box(0); } -lean_ctor_set(x_114, 0, x_108); -lean_ctor_set(x_114, 1, x_109); -lean_ctor_set(x_114, 2, x_110); -lean_ctor_set(x_114, 3, x_111); -lean_ctor_set_uint8(x_114, sizeof(void*)*4, x_113); -x_115 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_115, 0, x_107); -lean_ctor_set(x_115, 1, x_34); -lean_ctor_set(x_115, 2, x_35); -lean_ctor_set(x_115, 3, x_36); -lean_ctor_set_uint8(x_115, sizeof(void*)*4, x_113); -x_116 = 0; -lean_ctor_set(x_1, 3, x_115); -lean_ctor_set(x_1, 2, x_106); -lean_ctor_set(x_1, 1, x_105); -lean_ctor_set(x_1, 0, x_114); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_116); -return x_1; +x_443 = 1; +if (lean_is_scalar(x_442)) { + x_444 = lean_alloc_ctor(1, 4, 1); +} else { + x_444 = x_442; +} +lean_ctor_set(x_444, 0, x_342); +lean_ctor_set(x_444, 1, x_343); +lean_ctor_set(x_444, 2, x_344); +lean_ctor_set(x_444, 3, x_426); +lean_ctor_set_uint8(x_444, sizeof(void*)*4, x_443); +if (lean_is_scalar(x_437)) { + x_445 = lean_alloc_ctor(1, 4, 1); +} else { + x_445 = x_437; } +lean_ctor_set(x_445, 0, x_438); +lean_ctor_set(x_445, 1, x_439); +lean_ctor_set(x_445, 2, x_440); +lean_ctor_set(x_445, 3, x_441); +lean_ctor_set_uint8(x_445, sizeof(void*)*4, x_443); +x_446 = 0; +x_447 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_447, 0, x_444); +lean_ctor_set(x_447, 1, x_435); +lean_ctor_set(x_447, 2, x_436); +lean_ctor_set(x_447, 3, x_445); +lean_ctor_set_uint8(x_447, sizeof(void*)*4, x_446); +return x_447; } else { -lean_object* x_117; -x_117 = lean_ctor_get(x_38, 3); -lean_inc(x_117); -if (lean_obj_tag(x_117) == 0) -{ -uint8_t x_118; -lean_free_object(x_1); -x_118 = !lean_is_exclusive(x_40); -if (x_118 == 0) -{ -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_119 = lean_ctor_get(x_40, 3); -lean_dec(x_119); -x_120 = lean_ctor_get(x_40, 2); -lean_dec(x_120); -x_121 = lean_ctor_get(x_40, 1); -lean_dec(x_121); -x_122 = lean_ctor_get(x_40, 0); -lean_dec(x_122); -x_123 = 1; -lean_ctor_set(x_40, 3, x_36); -lean_ctor_set(x_40, 2, x_35); -lean_ctor_set(x_40, 1, x_34); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_123); -return x_40; +lean_object* x_448; uint8_t x_449; lean_object* x_450; +if (lean_is_exclusive(x_427)) { + lean_ctor_release(x_427, 0); + lean_ctor_release(x_427, 1); + lean_ctor_release(x_427, 2); + lean_ctor_release(x_427, 3); + x_448 = x_427; +} else { + lean_dec_ref(x_427); + x_448 = lean_box(0); +} +x_449 = 1; +if (lean_is_scalar(x_448)) { + x_450 = lean_alloc_ctor(1, 4, 1); +} else { + x_450 = x_448; +} +lean_ctor_set(x_450, 0, x_342); +lean_ctor_set(x_450, 1, x_343); +lean_ctor_set(x_450, 2, x_344); +lean_ctor_set(x_450, 3, x_424); +lean_ctor_set_uint8(x_450, sizeof(void*)*4, x_449); +return x_450; } -else -{ -uint8_t x_124; lean_object* x_125; -lean_dec(x_40); -x_124 = 1; -x_125 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_125, 0, x_38); -lean_ctor_set(x_125, 1, x_34); -lean_ctor_set(x_125, 2, x_35); -lean_ctor_set(x_125, 3, x_36); -lean_ctor_set_uint8(x_125, sizeof(void*)*4, x_124); -return x_125; } } else { -uint8_t x_126; -x_126 = lean_ctor_get_uint8(x_117, sizeof(void*)*4); -if (x_126 == 0) -{ -uint8_t x_127; -lean_free_object(x_1); -x_127 = !lean_is_exclusive(x_38); -if (x_127 == 0) -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_128 = lean_ctor_get(x_38, 1); -x_129 = lean_ctor_get(x_38, 2); -x_130 = lean_ctor_get(x_38, 3); -lean_dec(x_130); -x_131 = lean_ctor_get(x_38, 0); -lean_dec(x_131); -x_132 = !lean_is_exclusive(x_117); -if (x_132 == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; -x_133 = lean_ctor_get(x_117, 0); -x_134 = lean_ctor_get(x_117, 1); -x_135 = lean_ctor_get(x_117, 2); -x_136 = lean_ctor_get(x_117, 3); -x_137 = 1; -lean_inc(x_40); -lean_ctor_set(x_117, 3, x_133); -lean_ctor_set(x_117, 2, x_129); -lean_ctor_set(x_117, 1, x_128); -lean_ctor_set(x_117, 0, x_40); -x_138 = !lean_is_exclusive(x_40); -if (x_138 == 0) +uint8_t x_451; +x_451 = lean_ctor_get_uint8(x_426, sizeof(void*)*4); +if (x_451 == 0) { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; -x_139 = lean_ctor_get(x_40, 3); -lean_dec(x_139); -x_140 = lean_ctor_get(x_40, 2); -lean_dec(x_140); -x_141 = lean_ctor_get(x_40, 1); -lean_dec(x_141); -x_142 = lean_ctor_get(x_40, 0); -lean_dec(x_142); -lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); -lean_ctor_set(x_40, 3, x_36); -lean_ctor_set(x_40, 2, x_35); -lean_ctor_set(x_40, 1, x_34); -lean_ctor_set(x_40, 0, x_136); -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_137); -x_143 = 0; -lean_ctor_set(x_38, 3, x_40); -lean_ctor_set(x_38, 2, x_135); -lean_ctor_set(x_38, 1, x_134); -lean_ctor_set(x_38, 0, x_117); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_143); -return x_38; +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; uint8_t x_461; lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; +x_452 = lean_ctor_get(x_424, 1); +lean_inc(x_452); +x_453 = lean_ctor_get(x_424, 2); +lean_inc(x_453); +x_454 = lean_ctor_get(x_424, 3); +lean_inc(x_454); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_455 = x_424; +} else { + lean_dec_ref(x_424); + x_455 = lean_box(0); } -else -{ -lean_object* x_144; uint8_t x_145; -lean_dec(x_40); -lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); -x_144 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_144, 0, x_136); -lean_ctor_set(x_144, 1, x_34); -lean_ctor_set(x_144, 2, x_35); -lean_ctor_set(x_144, 3, x_36); -lean_ctor_set_uint8(x_144, sizeof(void*)*4, x_137); -x_145 = 0; -lean_ctor_set(x_38, 3, x_144); -lean_ctor_set(x_38, 2, x_135); -lean_ctor_set(x_38, 1, x_134); -lean_ctor_set(x_38, 0, x_117); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_145); -return x_38; +x_456 = lean_ctor_get(x_426, 0); +lean_inc(x_456); +x_457 = lean_ctor_get(x_426, 1); +lean_inc(x_457); +x_458 = lean_ctor_get(x_426, 2); +lean_inc(x_458); +x_459 = lean_ctor_get(x_426, 3); +lean_inc(x_459); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_460 = x_426; +} else { + lean_dec_ref(x_426); + x_460 = lean_box(0); +} +x_461 = 1; +if (lean_is_scalar(x_460)) { + x_462 = lean_alloc_ctor(1, 4, 1); +} else { + x_462 = x_460; +} +lean_ctor_set(x_462, 0, x_342); +lean_ctor_set(x_462, 1, x_343); +lean_ctor_set(x_462, 2, x_344); +lean_ctor_set(x_462, 3, x_456); +lean_ctor_set_uint8(x_462, sizeof(void*)*4, x_461); +if (lean_is_scalar(x_455)) { + x_463 = lean_alloc_ctor(1, 4, 1); +} else { + x_463 = x_455; } +lean_ctor_set(x_463, 0, x_459); +lean_ctor_set(x_463, 1, x_452); +lean_ctor_set(x_463, 2, x_453); +lean_ctor_set(x_463, 3, x_454); +lean_ctor_set_uint8(x_463, sizeof(void*)*4, x_461); +x_464 = 0; +x_465 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_465, 0, x_462); +lean_ctor_set(x_465, 1, x_457); +lean_ctor_set(x_465, 2, x_458); +lean_ctor_set(x_465, 3, x_463); +lean_ctor_set_uint8(x_465, sizeof(void*)*4, x_464); +return x_465; } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; -x_146 = lean_ctor_get(x_117, 0); -x_147 = lean_ctor_get(x_117, 1); -x_148 = lean_ctor_get(x_117, 2); -x_149 = lean_ctor_get(x_117, 3); -lean_inc(x_149); -lean_inc(x_148); -lean_inc(x_147); -lean_inc(x_146); -lean_dec(x_117); -x_150 = 1; -lean_inc(x_40); -x_151 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_151, 0, x_40); -lean_ctor_set(x_151, 1, x_128); -lean_ctor_set(x_151, 2, x_129); -lean_ctor_set(x_151, 3, x_146); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_152 = x_40; +lean_object* x_466; +x_466 = lean_ctor_get(x_424, 3); +lean_inc(x_466); +if (lean_obj_tag(x_466) == 0) +{ +lean_object* x_467; uint8_t x_468; lean_object* x_469; +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_467 = x_426; } else { - lean_dec_ref(x_40); - x_152 = lean_box(0); + lean_dec_ref(x_426); + x_467 = lean_box(0); } -lean_ctor_set_uint8(x_151, sizeof(void*)*4, x_150); -if (lean_is_scalar(x_152)) { - x_153 = lean_alloc_ctor(1, 4, 1); +x_468 = 1; +if (lean_is_scalar(x_467)) { + x_469 = lean_alloc_ctor(1, 4, 1); } else { - x_153 = x_152; -} -lean_ctor_set(x_153, 0, x_149); -lean_ctor_set(x_153, 1, x_34); -lean_ctor_set(x_153, 2, x_35); -lean_ctor_set(x_153, 3, x_36); -lean_ctor_set_uint8(x_153, sizeof(void*)*4, x_150); -x_154 = 0; -lean_ctor_set(x_38, 3, x_153); -lean_ctor_set(x_38, 2, x_148); -lean_ctor_set(x_38, 1, x_147); -lean_ctor_set(x_38, 0, x_151); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_154); -return x_38; + x_469 = x_467; } +lean_ctor_set(x_469, 0, x_342); +lean_ctor_set(x_469, 1, x_343); +lean_ctor_set(x_469, 2, x_344); +lean_ctor_set(x_469, 3, x_424); +lean_ctor_set_uint8(x_469, sizeof(void*)*4, x_468); +return x_469; } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; -x_155 = lean_ctor_get(x_38, 1); -x_156 = lean_ctor_get(x_38, 2); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_38); -x_157 = lean_ctor_get(x_117, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_117, 1); -lean_inc(x_158); -x_159 = lean_ctor_get(x_117, 2); -lean_inc(x_159); -x_160 = lean_ctor_get(x_117, 3); -lean_inc(x_160); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - lean_ctor_release(x_117, 2); - lean_ctor_release(x_117, 3); - x_161 = x_117; +uint8_t x_470; +x_470 = lean_ctor_get_uint8(x_466, sizeof(void*)*4); +if (x_470 == 0) +{ +lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; uint8_t x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; uint8_t x_483; lean_object* x_484; +x_471 = lean_ctor_get(x_424, 1); +lean_inc(x_471); +x_472 = lean_ctor_get(x_424, 2); +lean_inc(x_472); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_473 = x_424; } else { - lean_dec_ref(x_117); - x_161 = lean_box(0); + lean_dec_ref(x_424); + x_473 = lean_box(0); } -x_162 = 1; -lean_inc(x_40); -if (lean_is_scalar(x_161)) { - x_163 = lean_alloc_ctor(1, 4, 1); +x_474 = lean_ctor_get(x_466, 0); +lean_inc(x_474); +x_475 = lean_ctor_get(x_466, 1); +lean_inc(x_475); +x_476 = lean_ctor_get(x_466, 2); +lean_inc(x_476); +x_477 = lean_ctor_get(x_466, 3); +lean_inc(x_477); +if (lean_is_exclusive(x_466)) { + lean_ctor_release(x_466, 0); + lean_ctor_release(x_466, 1); + lean_ctor_release(x_466, 2); + lean_ctor_release(x_466, 3); + x_478 = x_466; } else { - x_163 = x_161; + lean_dec_ref(x_466); + x_478 = lean_box(0); } -lean_ctor_set(x_163, 0, x_40); -lean_ctor_set(x_163, 1, x_155); -lean_ctor_set(x_163, 2, x_156); -lean_ctor_set(x_163, 3, x_157); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_164 = x_40; +x_479 = 1; +lean_inc(x_426); +if (lean_is_scalar(x_478)) { + x_480 = lean_alloc_ctor(1, 4, 1); } else { - lean_dec_ref(x_40); - x_164 = lean_box(0); + x_480 = x_478; } -lean_ctor_set_uint8(x_163, sizeof(void*)*4, x_162); -if (lean_is_scalar(x_164)) { - x_165 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_480, 0, x_342); +lean_ctor_set(x_480, 1, x_343); +lean_ctor_set(x_480, 2, x_344); +lean_ctor_set(x_480, 3, x_426); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_481 = x_426; } else { - x_165 = x_164; + lean_dec_ref(x_426); + x_481 = lean_box(0); } -lean_ctor_set(x_165, 0, x_160); -lean_ctor_set(x_165, 1, x_34); -lean_ctor_set(x_165, 2, x_35); -lean_ctor_set(x_165, 3, x_36); -lean_ctor_set_uint8(x_165, sizeof(void*)*4, x_162); -x_166 = 0; -x_167 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_167, 0, x_163); -lean_ctor_set(x_167, 1, x_158); -lean_ctor_set(x_167, 2, x_159); -lean_ctor_set(x_167, 3, x_165); -lean_ctor_set_uint8(x_167, sizeof(void*)*4, x_166); -return x_167; +lean_ctor_set_uint8(x_480, sizeof(void*)*4, x_479); +if (lean_is_scalar(x_481)) { + x_482 = lean_alloc_ctor(1, 4, 1); +} else { + x_482 = x_481; } +lean_ctor_set(x_482, 0, x_474); +lean_ctor_set(x_482, 1, x_475); +lean_ctor_set(x_482, 2, x_476); +lean_ctor_set(x_482, 3, x_477); +lean_ctor_set_uint8(x_482, sizeof(void*)*4, x_479); +x_483 = 0; +if (lean_is_scalar(x_473)) { + x_484 = lean_alloc_ctor(1, 4, 1); +} else { + x_484 = x_473; } -else -{ -uint8_t x_168; -x_168 = !lean_is_exclusive(x_38); -if (x_168 == 0) -{ -lean_object* x_169; lean_object* x_170; uint8_t x_171; -x_169 = lean_ctor_get(x_38, 3); -lean_dec(x_169); -x_170 = lean_ctor_get(x_38, 0); -lean_dec(x_170); -x_171 = !lean_is_exclusive(x_40); -if (x_171 == 0) -{ -uint8_t x_172; -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_126); -x_172 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_172); -return x_1; +lean_ctor_set(x_484, 0, x_480); +lean_ctor_set(x_484, 1, x_471); +lean_ctor_set(x_484, 2, x_472); +lean_ctor_set(x_484, 3, x_482); +lean_ctor_set_uint8(x_484, sizeof(void*)*4, x_483); +return x_484; } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; -x_173 = lean_ctor_get(x_40, 0); -x_174 = lean_ctor_get(x_40, 1); -x_175 = lean_ctor_get(x_40, 2); -x_176 = lean_ctor_get(x_40, 3); -lean_inc(x_176); -lean_inc(x_175); -lean_inc(x_174); -lean_inc(x_173); -lean_dec(x_40); -x_177 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_177, 0, x_173); -lean_ctor_set(x_177, 1, x_174); -lean_ctor_set(x_177, 2, x_175); -lean_ctor_set(x_177, 3, x_176); -lean_ctor_set_uint8(x_177, sizeof(void*)*4, x_126); -lean_ctor_set(x_38, 0, x_177); -x_178 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_178); -return x_1; -} +lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; +x_485 = lean_ctor_get(x_424, 1); +lean_inc(x_485); +x_486 = lean_ctor_get(x_424, 2); +lean_inc(x_486); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_487 = x_424; +} else { + lean_dec_ref(x_424); + x_487 = lean_box(0); } -else -{ -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; -x_179 = lean_ctor_get(x_38, 1); -x_180 = lean_ctor_get(x_38, 2); -lean_inc(x_180); -lean_inc(x_179); -lean_dec(x_38); -x_181 = lean_ctor_get(x_40, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_40, 1); -lean_inc(x_182); -x_183 = lean_ctor_get(x_40, 2); -lean_inc(x_183); -x_184 = lean_ctor_get(x_40, 3); -lean_inc(x_184); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_185 = x_40; +x_488 = lean_ctor_get(x_426, 0); +lean_inc(x_488); +x_489 = lean_ctor_get(x_426, 1); +lean_inc(x_489); +x_490 = lean_ctor_get(x_426, 2); +lean_inc(x_490); +x_491 = lean_ctor_get(x_426, 3); +lean_inc(x_491); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_492 = x_426; } else { - lean_dec_ref(x_40); - x_185 = lean_box(0); + lean_dec_ref(x_426); + x_492 = lean_box(0); } -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 4, 1); +if (lean_is_scalar(x_492)) { + x_493 = lean_alloc_ctor(1, 4, 1); } else { - x_186 = x_185; + x_493 = x_492; } -lean_ctor_set(x_186, 0, x_181); -lean_ctor_set(x_186, 1, x_182); -lean_ctor_set(x_186, 2, x_183); -lean_ctor_set(x_186, 3, x_184); -lean_ctor_set_uint8(x_186, sizeof(void*)*4, x_126); -x_187 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_179); -lean_ctor_set(x_187, 2, x_180); -lean_ctor_set(x_187, 3, x_117); -lean_ctor_set_uint8(x_187, sizeof(void*)*4, x_39); -x_188 = 1; -lean_ctor_set(x_1, 0, x_187); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_188); -return x_1; +lean_ctor_set(x_493, 0, x_488); +lean_ctor_set(x_493, 1, x_489); +lean_ctor_set(x_493, 2, x_490); +lean_ctor_set(x_493, 3, x_491); +lean_ctor_set_uint8(x_493, sizeof(void*)*4, x_470); +if (lean_is_scalar(x_487)) { + x_494 = lean_alloc_ctor(1, 4, 1); +} else { + x_494 = x_487; } +lean_ctor_set(x_494, 0, x_493); +lean_ctor_set(x_494, 1, x_485); +lean_ctor_set(x_494, 2, x_486); +lean_ctor_set(x_494, 3, x_466); +lean_ctor_set_uint8(x_494, sizeof(void*)*4, x_425); +x_495 = 1; +x_496 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_496, 0, x_342); +lean_ctor_set(x_496, 1, x_343); +lean_ctor_set(x_496, 2, x_344); +lean_ctor_set(x_496, 3, x_494); +lean_ctor_set_uint8(x_496, sizeof(void*)*4, x_495); +return x_496; } } } @@ -11906,2492 +16658,2566 @@ return x_1; } else { -uint8_t x_189; -x_189 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_189); -return x_1; +uint8_t x_497; lean_object* x_498; +x_497 = 1; +x_498 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_498, 0, x_342); +lean_ctor_set(x_498, 1, x_343); +lean_ctor_set(x_498, 2, x_344); +lean_ctor_set(x_498, 3, x_424); +lean_ctor_set_uint8(x_498, sizeof(void*)*4, x_497); +return x_498; +} +} } } -case 1: -{ -uint8_t x_190; -lean_dec(x_35); -lean_dec(x_34); -x_190 = 1; -lean_ctor_set(x_1, 2, x_3); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_190); -return x_1; } -default: +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_191; uint8_t x_192; -x_191 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_36, x_2, x_3); -x_192 = lean_ctor_get_uint8(x_191, sizeof(void*)*4); -if (x_192 == 0) +uint8_t x_4; +x_4 = l_Lean_RBNode_isRed___rarg(x_1); +if (x_4 == 0) { -lean_object* x_193; -x_193 = lean_ctor_get(x_191, 0); -lean_inc(x_193); -if (lean_obj_tag(x_193) == 0) +lean_object* x_5; +x_5 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); +return x_5; +} +else { -lean_object* x_194; -x_194 = lean_ctor_get(x_191, 3); -lean_inc(x_194); -if (lean_obj_tag(x_194) == 0) +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); +x_7 = l_Lean_RBNode_setBlack___rarg(x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -uint8_t x_195; -x_195 = !lean_is_exclusive(x_191); -if (x_195 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_196; lean_object* x_197; uint8_t x_198; -x_196 = lean_ctor_get(x_191, 3); -lean_dec(x_196); -x_197 = lean_ctor_get(x_191, 0); -lean_dec(x_197); -lean_ctor_set(x_191, 0, x_194); -x_198 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_198); -return x_1; +lean_dec(x_5); +return x_6; } else { -lean_object* x_199; lean_object* x_200; lean_object* x_201; uint8_t x_202; -x_199 = lean_ctor_get(x_191, 1); -x_200 = lean_ctor_get(x_191, 2); -lean_inc(x_200); -lean_inc(x_199); -lean_dec(x_191); -x_201 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_201, 0, x_194); -lean_ctor_set(x_201, 1, x_199); -lean_ctor_set(x_201, 2, x_200); -lean_ctor_set(x_201, 3, x_194); -lean_ctor_set_uint8(x_201, sizeof(void*)*4, x_192); -x_202 = 1; -lean_ctor_set(x_1, 3, x_201); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_202); -return x_1; +lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_9); +x_12 = lean_uint64_to_usize(x_11); +x_13 = 1; +x_14 = lean_usize_sub(x_1, x_13); +x_15 = 5; +x_16 = lean_usize_mul(x_15, x_14); +x_17 = lean_usize_shift_right(x_12, x_16); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_5, x_18); +lean_dec(x_5); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_6, x_17, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_19; +x_6 = x_20; +goto _start; } } -else -{ -uint8_t x_203; -x_203 = lean_ctor_get_uint8(x_194, sizeof(void*)*4); -if (x_203 == 0) +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_204; -x_204 = !lean_is_exclusive(x_191); -if (x_204 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; -x_205 = lean_ctor_get(x_191, 1); -x_206 = lean_ctor_get(x_191, 2); -x_207 = lean_ctor_get(x_191, 3); -lean_dec(x_207); -x_208 = lean_ctor_get(x_191, 0); -lean_dec(x_208); -x_209 = !lean_is_exclusive(x_194); -if (x_209 == 0) +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) { -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; uint8_t x_215; -x_210 = lean_ctor_get(x_194, 0); -x_211 = lean_ctor_get(x_194, 1); -x_212 = lean_ctor_get(x_194, 2); -x_213 = lean_ctor_get(x_194, 3); -x_214 = 1; -lean_ctor_set(x_194, 3, x_193); -lean_ctor_set(x_194, 2, x_35); -lean_ctor_set(x_194, 1, x_34); -lean_ctor_set(x_194, 0, x_33); -lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_214); -lean_ctor_set(x_191, 3, x_213); -lean_ctor_set(x_191, 2, x_212); -lean_ctor_set(x_191, 1, x_211); -lean_ctor_set(x_191, 0, x_210); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_214); -x_215 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_206); -lean_ctor_set(x_1, 1, x_205); -lean_ctor_set(x_1, 0, x_194); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_215); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); return x_1; } else { -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; uint8_t x_222; -x_216 = lean_ctor_get(x_194, 0); -x_217 = lean_ctor_get(x_194, 1); -x_218 = lean_ctor_get(x_194, 2); -x_219 = lean_ctor_get(x_194, 3); -lean_inc(x_219); -lean_inc(x_218); -lean_inc(x_217); -lean_inc(x_216); -lean_dec(x_194); -x_220 = 1; -x_221 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_221, 0, x_33); -lean_ctor_set(x_221, 1, x_34); -lean_ctor_set(x_221, 2, x_35); -lean_ctor_set(x_221, 3, x_193); -lean_ctor_set_uint8(x_221, sizeof(void*)*4, x_220); -lean_ctor_set(x_191, 3, x_219); -lean_ctor_set(x_191, 2, x_218); -lean_ctor_set(x_191, 1, x_217); -lean_ctor_set(x_191, 0, x_216); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_220); -x_222 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_206); -lean_ctor_set(x_1, 1, x_205); -lean_ctor_set(x_1, 0, x_221); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_222); -return x_1; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; uint8_t x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; -x_223 = lean_ctor_get(x_191, 1); -x_224 = lean_ctor_get(x_191, 2); -lean_inc(x_224); -lean_inc(x_223); -lean_dec(x_191); -x_225 = lean_ctor_get(x_194, 0); -lean_inc(x_225); -x_226 = lean_ctor_get(x_194, 1); -lean_inc(x_226); -x_227 = lean_ctor_get(x_194, 2); -lean_inc(x_227); -x_228 = lean_ctor_get(x_194, 3); -lean_inc(x_228); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - lean_ctor_release(x_194, 2); - lean_ctor_release(x_194, 3); - x_229 = x_194; -} else { - lean_dec_ref(x_194); - x_229 = lean_box(0); -} -x_230 = 1; -if (lean_is_scalar(x_229)) { - x_231 = lean_alloc_ctor(1, 4, 1); -} else { - x_231 = x_229; -} -lean_ctor_set(x_231, 0, x_33); -lean_ctor_set(x_231, 1, x_34); -lean_ctor_set(x_231, 2, x_35); -lean_ctor_set(x_231, 3, x_193); -lean_ctor_set_uint8(x_231, sizeof(void*)*4, x_230); -x_232 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_232, 0, x_225); -lean_ctor_set(x_232, 1, x_226); -lean_ctor_set(x_232, 2, x_227); -lean_ctor_set(x_232, 3, x_228); -lean_ctor_set_uint8(x_232, sizeof(void*)*4, x_230); -x_233 = 0; -lean_ctor_set(x_1, 3, x_232); -lean_ctor_set(x_1, 2, x_224); -lean_ctor_set(x_1, 1, x_223); -lean_ctor_set(x_1, 0, x_231); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_233); -return x_1; -} +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_6); +lean_dec(x_5); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_2); +x_2 = x_20; +goto _start; } else { -uint8_t x_234; -lean_free_object(x_1); -x_234 = !lean_is_exclusive(x_194); -if (x_234 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 0) { -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; -x_235 = lean_ctor_get(x_194, 3); -lean_dec(x_235); -x_236 = lean_ctor_get(x_194, 2); -lean_dec(x_236); -x_237 = lean_ctor_get(x_194, 1); -lean_dec(x_237); -x_238 = lean_ctor_get(x_194, 0); -lean_dec(x_238); -x_239 = 1; -lean_ctor_set(x_194, 3, x_191); -lean_ctor_set(x_194, 2, x_35); -lean_ctor_set(x_194, 1, x_34); -lean_ctor_set(x_194, 0, x_33); -lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_239); -return x_194; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_1, 0); +lean_dec(x_24); +x_25 = lean_array_fset(x_5, x_2, x_3); +x_26 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); +return x_1; } else { -uint8_t x_240; lean_object* x_241; -lean_dec(x_194); -x_240 = 1; -x_241 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_241, 0, x_33); -lean_ctor_set(x_241, 1, x_34); -lean_ctor_set(x_241, 2, x_35); -lean_ctor_set(x_241, 3, x_191); -lean_ctor_set_uint8(x_241, sizeof(void*)*4, x_240); -return x_241; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; +} else { -uint8_t x_242; -x_242 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); -if (x_242 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -uint8_t x_243; -x_243 = !lean_is_exclusive(x_191); -if (x_243 == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_244; uint8_t x_245; -x_244 = lean_ctor_get(x_191, 0); -lean_dec(x_244); -x_245 = !lean_is_exclusive(x_193); -if (x_245 == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_19); +if (x_21 == 0) { -lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; uint8_t x_251; -x_246 = lean_ctor_get(x_193, 0); -x_247 = lean_ctor_get(x_193, 1); -x_248 = lean_ctor_get(x_193, 2); -x_249 = lean_ctor_get(x_193, 3); -x_250 = 1; -lean_ctor_set(x_193, 3, x_246); -lean_ctor_set(x_193, 2, x_35); -lean_ctor_set(x_193, 1, x_34); -lean_ctor_set(x_193, 0, x_33); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_250); -lean_ctor_set(x_191, 0, x_249); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_250); -x_251 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_248); -lean_ctor_set(x_1, 1, x_247); -lean_ctor_set(x_1, 0, x_193); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_251); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_array_fset(x_17, x_12, x_23); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_24); return x_1; } else { -lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; uint8_t x_256; lean_object* x_257; uint8_t x_258; -x_252 = lean_ctor_get(x_193, 0); -x_253 = lean_ctor_get(x_193, 1); -x_254 = lean_ctor_get(x_193, 2); -x_255 = lean_ctor_get(x_193, 3); -lean_inc(x_255); -lean_inc(x_254); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_193); -x_256 = 1; -x_257 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_257, 0, x_33); -lean_ctor_set(x_257, 1, x_34); -lean_ctor_set(x_257, 2, x_35); -lean_ctor_set(x_257, 3, x_252); -lean_ctor_set_uint8(x_257, sizeof(void*)*4, x_256); -lean_ctor_set(x_191, 0, x_255); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_256); -x_258 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_254); -lean_ctor_set(x_1, 1, x_253); -lean_ctor_set(x_1, 0, x_257); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_258); +lean_object* x_25; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_25 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_25); return x_1; } } else { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint8_t x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; -x_259 = lean_ctor_get(x_191, 1); -x_260 = lean_ctor_get(x_191, 2); -x_261 = lean_ctor_get(x_191, 3); -lean_inc(x_261); -lean_inc(x_260); -lean_inc(x_259); -lean_dec(x_191); -x_262 = lean_ctor_get(x_193, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_193, 1); -lean_inc(x_263); -x_264 = lean_ctor_get(x_193, 2); -lean_inc(x_264); -x_265 = lean_ctor_get(x_193, 3); -lean_inc(x_265); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_266 = x_193; -} else { - lean_dec_ref(x_193); - x_266 = lean_box(0); +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = lean_array_fset(x_17, x_12, x_30); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_31); +return x_1; } -x_267 = 1; -if (lean_is_scalar(x_266)) { - x_268 = lean_alloc_ctor(1, 4, 1); -} else { - x_268 = x_266; +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_27); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_5); +x_33 = lean_array_fset(x_17, x_12, x_32); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_33); +return x_1; } -lean_ctor_set(x_268, 0, x_33); -lean_ctor_set(x_268, 1, x_34); -lean_ctor_set(x_268, 2, x_35); -lean_ctor_set(x_268, 3, x_262); -lean_ctor_set_uint8(x_268, sizeof(void*)*4, x_267); -x_269 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_269, 0, x_265); -lean_ctor_set(x_269, 1, x_259); -lean_ctor_set(x_269, 2, x_260); -lean_ctor_set(x_269, 3, x_261); -lean_ctor_set_uint8(x_269, sizeof(void*)*4, x_267); -x_270 = 0; -lean_ctor_set(x_1, 3, x_269); -lean_ctor_set(x_1, 2, x_264); -lean_ctor_set(x_1, 1, x_263); -lean_ctor_set(x_1, 0, x_268); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_270); +} +} +case 1: +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) +{ +lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_usize_shift_right(x_2, x_9); +x_37 = lean_usize_add(x_3, x_8); +x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_35, x_36, x_37, x_4, x_5); +lean_ctor_set(x_15, 0, x_38); +x_39 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); return x_1; } -} else { -lean_object* x_271; -x_271 = lean_ctor_get(x_191, 3); -lean_inc(x_271); -if (lean_obj_tag(x_271) == 0) -{ -uint8_t x_272; -lean_free_object(x_1); -x_272 = !lean_is_exclusive(x_193); -if (x_272 == 0) -{ -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; uint8_t x_277; -x_273 = lean_ctor_get(x_193, 3); -lean_dec(x_273); -x_274 = lean_ctor_get(x_193, 2); -lean_dec(x_274); -x_275 = lean_ctor_get(x_193, 1); -lean_dec(x_275); -x_276 = lean_ctor_get(x_193, 0); -lean_dec(x_276); -x_277 = 1; -lean_ctor_set(x_193, 3, x_191); -lean_ctor_set(x_193, 2, x_35); -lean_ctor_set(x_193, 1, x_34); -lean_ctor_set(x_193, 0, x_33); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_277); -return x_193; +lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_15, 0); +lean_inc(x_40); +lean_dec(x_15); +x_41 = lean_usize_shift_right(x_2, x_9); +x_42 = lean_usize_add(x_3, x_8); +x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_40, x_41, x_42, x_4, x_5); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_array_fset(x_17, x_12, x_44); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_45); +return x_1; } -else +} +default: { -uint8_t x_278; lean_object* x_279; -lean_dec(x_193); -x_278 = 1; -x_279 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_279, 0, x_33); -lean_ctor_set(x_279, 1, x_34); -lean_ctor_set(x_279, 2, x_35); -lean_ctor_set(x_279, 3, x_191); -lean_ctor_set_uint8(x_279, sizeof(void*)*4, x_278); -return x_279; +lean_object* x_46; lean_object* x_47; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_4); +lean_ctor_set(x_46, 1, x_5); +x_47 = lean_array_fset(x_17, x_12, x_46); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_47); +return x_1; +} +} } } else { -uint8_t x_280; -x_280 = lean_ctor_get_uint8(x_271, sizeof(void*)*4); -if (x_280 == 0) +lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +lean_dec(x_1); +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_52 = lean_usize_land(x_2, x_51); +x_53 = lean_usize_to_nat(x_52); +x_54 = lean_array_get_size(x_48); +x_55 = lean_nat_dec_lt(x_53, x_54); +lean_dec(x_54); +if (x_55 == 0) { -uint8_t x_281; -lean_free_object(x_1); -x_281 = !lean_is_exclusive(x_191); -if (x_281 == 0) +lean_object* x_56; +lean_dec(x_53); +lean_dec(x_5); +lean_dec(x_4); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; +} +else { -lean_object* x_282; lean_object* x_283; uint8_t x_284; -x_282 = lean_ctor_get(x_191, 3); -lean_dec(x_282); -x_283 = lean_ctor_get(x_191, 0); -lean_dec(x_283); -x_284 = !lean_is_exclusive(x_271); -if (x_284 == 0) +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_array_fget(x_48, x_53); +x_58 = lean_box(0); +x_59 = lean_array_fset(x_48, x_53, x_58); +switch (lean_obj_tag(x_57)) { +case 0: { -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; uint8_t x_290; -x_285 = lean_ctor_get(x_271, 0); -x_286 = lean_ctor_get(x_271, 1); -x_287 = lean_ctor_get(x_271, 2); -x_288 = lean_ctor_get(x_271, 3); -x_289 = 1; -lean_inc(x_193); -lean_ctor_set(x_271, 3, x_193); -lean_ctor_set(x_271, 2, x_35); -lean_ctor_set(x_271, 1, x_34); -lean_ctor_set(x_271, 0, x_33); -x_290 = !lean_is_exclusive(x_193); -if (x_290 == 0) +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_62 = x_57; +} else { + lean_dec_ref(x_57); + x_62 = lean_box(0); +} +x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); +if (x_63 == 0) { -lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; -x_291 = lean_ctor_get(x_193, 3); -lean_dec(x_291); -x_292 = lean_ctor_get(x_193, 2); -lean_dec(x_292); -x_293 = lean_ctor_get(x_193, 1); -lean_dec(x_293); -x_294 = lean_ctor_get(x_193, 0); -lean_dec(x_294); -lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); -lean_ctor_set(x_193, 3, x_288); -lean_ctor_set(x_193, 2, x_287); -lean_ctor_set(x_193, 1, x_286); -lean_ctor_set(x_193, 0, x_285); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_289); -x_295 = 0; -lean_ctor_set(x_191, 3, x_193); -lean_ctor_set(x_191, 0, x_271); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_295); -return x_191; +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_66 = lean_array_fset(x_59, x_53, x_65); +lean_dec(x_53); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } else { -lean_object* x_296; uint8_t x_297; -lean_dec(x_193); -lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); -x_296 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_296, 0, x_285); -lean_ctor_set(x_296, 1, x_286); -lean_ctor_set(x_296, 2, x_287); -lean_ctor_set(x_296, 3, x_288); -lean_ctor_set_uint8(x_296, sizeof(void*)*4, x_289); -x_297 = 0; -lean_ctor_set(x_191, 3, x_296); -lean_ctor_set(x_191, 0, x_271); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_297); -return x_191; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_61); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_62; +} +lean_ctor_set(x_68, 0, x_4); +lean_ctor_set(x_68, 1, x_5); +x_69 = lean_array_fset(x_59, x_53, x_68); +lean_dec(x_53); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; } } -else +case 1: { -lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; uint8_t x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; -x_298 = lean_ctor_get(x_271, 0); -x_299 = lean_ctor_get(x_271, 1); -x_300 = lean_ctor_get(x_271, 2); -x_301 = lean_ctor_get(x_271, 3); -lean_inc(x_301); -lean_inc(x_300); -lean_inc(x_299); -lean_inc(x_298); -lean_dec(x_271); -x_302 = 1; -lean_inc(x_193); -x_303 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_303, 0, x_33); -lean_ctor_set(x_303, 1, x_34); -lean_ctor_set(x_303, 2, x_35); -lean_ctor_set(x_303, 3, x_193); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_304 = x_193; +lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_71 = lean_ctor_get(x_57, 0); +lean_inc(x_71); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_72 = x_57; } else { - lean_dec_ref(x_193); - x_304 = lean_box(0); + lean_dec_ref(x_57); + x_72 = lean_box(0); +} +x_73 = lean_usize_shift_right(x_2, x_50); +x_74 = lean_usize_add(x_3, x_49); +x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_71, x_73, x_74, x_4, x_5); +if (lean_is_scalar(x_72)) { + x_76 = lean_alloc_ctor(1, 1, 0); +} else { + x_76 = x_72; +} +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_array_fset(x_59, x_53, x_76); +lean_dec(x_53); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; +} +default: +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_4); +lean_ctor_set(x_79, 1, x_5); +x_80 = lean_array_fset(x_59, x_53, x_79); +lean_dec(x_53); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} } -lean_ctor_set_uint8(x_303, sizeof(void*)*4, x_302); -if (lean_is_scalar(x_304)) { - x_305 = lean_alloc_ctor(1, 4, 1); -} else { - x_305 = x_304; } -lean_ctor_set(x_305, 0, x_298); -lean_ctor_set(x_305, 1, x_299); -lean_ctor_set(x_305, 2, x_300); -lean_ctor_set(x_305, 3, x_301); -lean_ctor_set_uint8(x_305, sizeof(void*)*4, x_302); -x_306 = 0; -lean_ctor_set(x_191, 3, x_305); -lean_ctor_set(x_191, 0, x_303); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_306); -return x_191; } } else { -lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; uint8_t x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; uint8_t x_318; lean_object* x_319; -x_307 = lean_ctor_get(x_191, 1); -x_308 = lean_ctor_get(x_191, 2); -lean_inc(x_308); -lean_inc(x_307); -lean_dec(x_191); -x_309 = lean_ctor_get(x_271, 0); -lean_inc(x_309); -x_310 = lean_ctor_get(x_271, 1); -lean_inc(x_310); -x_311 = lean_ctor_get(x_271, 2); -lean_inc(x_311); -x_312 = lean_ctor_get(x_271, 3); -lean_inc(x_312); -if (lean_is_exclusive(x_271)) { - lean_ctor_release(x_271, 0); - lean_ctor_release(x_271, 1); - lean_ctor_release(x_271, 2); - lean_ctor_release(x_271, 3); - x_313 = x_271; -} else { - lean_dec_ref(x_271); - x_313 = lean_box(0); -} -x_314 = 1; -lean_inc(x_193); -if (lean_is_scalar(x_313)) { - x_315 = lean_alloc_ctor(1, 4, 1); -} else { - x_315 = x_313; +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; +x_83 = lean_unsigned_to_nat(0u); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(x_1, x_83, x_4, x_5); +x_85 = 7; +x_86 = lean_usize_dec_le(x_85, x_3); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_dec_lt(x_87, x_88); +lean_dec(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_84, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 1); +lean_inc(x_91); +lean_dec(x_84); +x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } -lean_ctor_set(x_315, 0, x_33); -lean_ctor_set(x_315, 1, x_34); -lean_ctor_set(x_315, 2, x_35); -lean_ctor_set(x_315, 3, x_193); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_316 = x_193; -} else { - lean_dec_ref(x_193); - x_316 = lean_box(0); +else +{ +return x_84; } -lean_ctor_set_uint8(x_315, sizeof(void*)*4, x_314); -if (lean_is_scalar(x_316)) { - x_317 = lean_alloc_ctor(1, 4, 1); -} else { - x_317 = x_316; } -lean_ctor_set(x_317, 0, x_309); -lean_ctor_set(x_317, 1, x_310); -lean_ctor_set(x_317, 2, x_311); -lean_ctor_set(x_317, 3, x_312); -lean_ctor_set_uint8(x_317, sizeof(void*)*4, x_314); -x_318 = 0; -x_319 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_319, 0, x_315); -lean_ctor_set(x_319, 1, x_307); -lean_ctor_set(x_319, 2, x_308); -lean_ctor_set(x_319, 3, x_317); -lean_ctor_set_uint8(x_319, sizeof(void*)*4, x_318); -return x_319; +else +{ +return x_84; } } else { -uint8_t x_320; -x_320 = !lean_is_exclusive(x_191); -if (x_320 == 0) +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_1); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_unsigned_to_nat(0u); +x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(x_96, x_97, x_4, x_5); +x_99 = 7; +x_100 = lean_usize_dec_le(x_99, x_3); +if (x_100 == 0) { -lean_object* x_321; lean_object* x_322; uint8_t x_323; -x_321 = lean_ctor_get(x_191, 3); -lean_dec(x_321); -x_322 = lean_ctor_get(x_191, 0); -lean_dec(x_322); -x_323 = !lean_is_exclusive(x_193); -if (x_323 == 0) +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_dec_lt(x_101, x_102); +lean_dec(x_101); +if (x_103 == 0) { -uint8_t x_324; -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_280); -x_324 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_324); -return x_1; +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; } else { -lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint8_t x_330; -x_325 = lean_ctor_get(x_193, 0); -x_326 = lean_ctor_get(x_193, 1); -x_327 = lean_ctor_get(x_193, 2); -x_328 = lean_ctor_get(x_193, 3); -lean_inc(x_328); -lean_inc(x_327); -lean_inc(x_326); -lean_inc(x_325); -lean_dec(x_193); -x_329 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_329, 0, x_325); -lean_ctor_set(x_329, 1, x_326); -lean_ctor_set(x_329, 2, x_327); -lean_ctor_set(x_329, 3, x_328); -lean_ctor_set_uint8(x_329, sizeof(void*)*4, x_280); -lean_ctor_set(x_191, 0, x_329); -x_330 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_330); -return x_1; +return x_98; } } else { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; uint8_t x_340; -x_331 = lean_ctor_get(x_191, 1); -x_332 = lean_ctor_get(x_191, 2); -lean_inc(x_332); -lean_inc(x_331); -lean_dec(x_191); -x_333 = lean_ctor_get(x_193, 0); -lean_inc(x_333); -x_334 = lean_ctor_get(x_193, 1); -lean_inc(x_334); -x_335 = lean_ctor_get(x_193, 2); -lean_inc(x_335); -x_336 = lean_ctor_get(x_193, 3); -lean_inc(x_336); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_337 = x_193; -} else { - lean_dec_ref(x_193); - x_337 = lean_box(0); +return x_98; } -if (lean_is_scalar(x_337)) { - x_338 = lean_alloc_ctor(1, 4, 1); -} else { - x_338 = x_337; } -lean_ctor_set(x_338, 0, x_333); -lean_ctor_set(x_338, 1, x_334); -lean_ctor_set(x_338, 2, x_335); -lean_ctor_set(x_338, 3, x_336); -lean_ctor_set_uint8(x_338, sizeof(void*)*4, x_280); -x_339 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_339, 0, x_338); -lean_ctor_set(x_339, 1, x_331); -lean_ctor_set(x_339, 2, x_332); -lean_ctor_set(x_339, 3, x_271); -lean_ctor_set_uint8(x_339, sizeof(void*)*4, x_192); -x_340 = 1; -lean_ctor_set(x_1, 3, x_339); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_340); -return x_1; } } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_1, x_5, x_6, x_2, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_box(0); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_array_fget(x_2, x_4); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = 5; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_8 = lean_usize_land(x_2, x_7); +x_9 = lean_usize_to_nat(x_8); +x_10 = lean_box(2); +x_11 = lean_array_get(x_10, x_5, x_9); +lean_dec(x_9); +lean_dec(x_5); +switch (lean_obj_tag(x_11)) { +case 0: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_12); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_13); +lean_free_object(x_1); +x_15 = lean_box(0); +return x_15; } else { -uint8_t x_341; -x_341 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_341); +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_13); return x_1; } } +case 1: +{ +lean_object* x_16; size_t x_17; +lean_free_object(x_1); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_usize_shift_right(x_2, x_6); +x_1 = x_16; +x_2 = x_17; +goto _start; +} +default: +{ +lean_object* x_19; +lean_free_object(x_1); +x_19 = lean_box(0); +return x_19; +} } } else { -lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; -x_342 = lean_ctor_get(x_1, 0); -x_343 = lean_ctor_get(x_1, 1); -x_344 = lean_ctor_get(x_1, 2); -x_345 = lean_ctor_get(x_1, 3); -lean_inc(x_345); -lean_inc(x_344); -lean_inc(x_343); -lean_inc(x_342); +lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); lean_dec(x_1); -x_346 = l_Lean_Expr_quickComp(x_2, x_343); -switch (x_346) { +x_21 = 5; +x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_23 = lean_usize_land(x_2, x_22); +x_24 = lean_usize_to_nat(x_23); +x_25 = lean_box(2); +x_26 = lean_array_get(x_25, x_20, x_24); +lean_dec(x_24); +lean_dec(x_20); +switch (lean_obj_tag(x_26)) { case 0: { -lean_object* x_347; uint8_t x_348; -x_347 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_342, x_2, x_3); -x_348 = lean_ctor_get_uint8(x_347, sizeof(void*)*4); -if (x_348 == 0) +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_27); +lean_dec(x_27); +if (x_29 == 0) { -lean_object* x_349; -x_349 = lean_ctor_get(x_347, 0); -lean_inc(x_349); -if (lean_obj_tag(x_349) == 0) +lean_object* x_30; +lean_dec(x_28); +x_30 = lean_box(0); +return x_30; +} +else { -lean_object* x_350; -x_350 = lean_ctor_get(x_347, 3); -lean_inc(x_350); -if (lean_obj_tag(x_350) == 0) +lean_object* x_31; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_28); +return x_31; +} +} +case 1: { -lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; uint8_t x_355; lean_object* x_356; -x_351 = lean_ctor_get(x_347, 1); -lean_inc(x_351); -x_352 = lean_ctor_get(x_347, 2); -lean_inc(x_352); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_353 = x_347; -} else { - lean_dec_ref(x_347); - x_353 = lean_box(0); +lean_object* x_32; size_t x_33; +x_32 = lean_ctor_get(x_26, 0); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_usize_shift_right(x_2, x_21); +x_1 = x_32; +x_2 = x_33; +goto _start; +} +default: +{ +lean_object* x_35; +x_35 = lean_box(0); +return x_35; +} } -if (lean_is_scalar(x_353)) { - x_354 = lean_alloc_ctor(1, 4, 1); -} else { - x_354 = x_353; } -lean_ctor_set(x_354, 0, x_350); -lean_ctor_set(x_354, 1, x_351); -lean_ctor_set(x_354, 2, x_352); -lean_ctor_set(x_354, 3, x_350); -lean_ctor_set_uint8(x_354, sizeof(void*)*4, x_348); -x_355 = 1; -x_356 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_356, 0, x_354); -lean_ctor_set(x_356, 1, x_343); -lean_ctor_set(x_356, 2, x_344); -lean_ctor_set(x_356, 3, x_345); -lean_ctor_set_uint8(x_356, sizeof(void*)*4, x_355); -return x_356; } else { -uint8_t x_357; -x_357 = lean_ctor_get_uint8(x_350, sizeof(void*)*4); -if (x_357 == 0) -{ -lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; uint8_t x_366; lean_object* x_367; lean_object* x_368; uint8_t x_369; lean_object* x_370; -x_358 = lean_ctor_get(x_347, 1); -lean_inc(x_358); -x_359 = lean_ctor_get(x_347, 2); -lean_inc(x_359); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_360 = x_347; -} else { - lean_dec_ref(x_347); - x_360 = lean_box(0); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_1, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(x_36, x_37, lean_box(0), x_38, x_3); +lean_dec(x_37); +lean_dec(x_36); +return x_39; } -x_361 = lean_ctor_get(x_350, 0); -lean_inc(x_361); -x_362 = lean_ctor_get(x_350, 1); -lean_inc(x_362); -x_363 = lean_ctor_get(x_350, 2); -lean_inc(x_363); -x_364 = lean_ctor_get(x_350, 3); -lean_inc(x_364); -if (lean_is_exclusive(x_350)) { - lean_ctor_release(x_350, 0); - lean_ctor_release(x_350, 1); - lean_ctor_release(x_350, 2); - lean_ctor_release(x_350, 3); - x_365 = x_350; -} else { - lean_dec_ref(x_350); - x_365 = lean_box(0); } -x_366 = 1; -if (lean_is_scalar(x_365)) { - x_367 = lean_alloc_ctor(1, 4, 1); -} else { - x_367 = x_365; } -lean_ctor_set(x_367, 0, x_349); -lean_ctor_set(x_367, 1, x_358); -lean_ctor_set(x_367, 2, x_359); -lean_ctor_set(x_367, 3, x_361); -lean_ctor_set_uint8(x_367, sizeof(void*)*4, x_366); -if (lean_is_scalar(x_360)) { - x_368 = lean_alloc_ctor(1, 4, 1); -} else { - x_368 = x_360; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(x_1, x_4, x_2); +return x_5; } -lean_ctor_set(x_368, 0, x_364); -lean_ctor_set(x_368, 1, x_343); -lean_ctor_set(x_368, 2, x_344); -lean_ctor_set(x_368, 3, x_345); -lean_ctor_set_uint8(x_368, sizeof(void*)*4, x_366); -x_369 = 0; -x_370 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_370, 0, x_367); -lean_ctor_set(x_370, 1, x_362); -lean_ctor_set(x_370, 2, x_363); -lean_ctor_set(x_370, 3, x_368); -lean_ctor_set_uint8(x_370, sizeof(void*)*4, x_369); -return x_370; +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; +x_12 = l_Lean_Meta_Grind_getRoot_x3f(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +lean_dec(x_1); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_12, 0, x_16); +return x_12; } else { -lean_object* x_371; uint8_t x_372; lean_object* x_373; -if (lean_is_exclusive(x_350)) { - lean_ctor_release(x_350, 0); - lean_ctor_release(x_350, 1); - lean_ctor_release(x_350, 2); - lean_ctor_release(x_350, 3); - x_371 = x_350; -} else { - lean_dec_ref(x_350); - x_371 = lean_box(0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; } -x_372 = 1; -if (lean_is_scalar(x_371)) { - x_373 = lean_alloc_ctor(1, 4, 1); -} else { - x_373 = x_371; } -lean_ctor_set(x_373, 0, x_347); -lean_ctor_set(x_373, 1, x_343); -lean_ctor_set(x_373, 2, x_344); -lean_ctor_set(x_373, 3, x_345); -lean_ctor_set_uint8(x_373, sizeof(void*)*4, x_372); -return x_373; +else +{ +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_68; lean_object* x_69; +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_ctor_get(x_13, 0); +lean_inc(x_21); +lean_dec(x_13); +x_22 = lean_st_ref_get(x_3, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 2); +lean_inc(x_25); +lean_dec(x_23); +x_68 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_25, x_21); +x_69 = lean_st_ref_take(x_3, x_24); +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = lean_box(0); +x_26 = x_72; +x_27 = x_70; +x_28 = x_71; +goto block_67; } +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_68, 0); +lean_inc(x_73); +lean_dec(x_68); +x_74 = lean_ctor_get(x_69, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_69, 1); +lean_inc(x_75); +lean_dec(x_69); +x_26 = x_73; +x_27 = x_74; +x_28 = x_75; +goto block_67; } +block_67: +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_30 = lean_ctor_get(x_27, 2); +x_31 = lean_box(0); +x_32 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_31); +x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_30, x_21, x_32); +lean_ctor_set(x_27, 2, x_33); +x_34 = lean_st_ref_set(x_3, x_27, x_28); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_34, 0); +lean_dec(x_36); +lean_ctor_set(x_34, 0, x_31); +return x_34; } else { -uint8_t x_374; -x_374 = lean_ctor_get_uint8(x_349, sizeof(void*)*4); -if (x_374 == 0) +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_31); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else { -lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; uint8_t x_384; lean_object* x_385; lean_object* x_386; uint8_t x_387; lean_object* x_388; -x_375 = lean_ctor_get(x_347, 1); -lean_inc(x_375); -x_376 = lean_ctor_get(x_347, 2); -lean_inc(x_376); -x_377 = lean_ctor_get(x_347, 3); -lean_inc(x_377); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_378 = x_347; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_39 = lean_ctor_get(x_27, 0); +x_40 = lean_ctor_get(x_27, 1); +x_41 = lean_ctor_get(x_27, 2); +x_42 = lean_ctor_get(x_27, 3); +x_43 = lean_ctor_get(x_27, 4); +x_44 = lean_ctor_get(x_27, 5); +x_45 = lean_ctor_get_uint8(x_27, sizeof(void*)*19); +x_46 = lean_ctor_get(x_27, 6); +x_47 = lean_ctor_get(x_27, 7); +x_48 = lean_ctor_get(x_27, 8); +x_49 = lean_ctor_get(x_27, 9); +x_50 = lean_ctor_get(x_27, 10); +x_51 = lean_ctor_get(x_27, 11); +x_52 = lean_ctor_get(x_27, 12); +x_53 = lean_ctor_get(x_27, 13); +x_54 = lean_ctor_get(x_27, 14); +x_55 = lean_ctor_get(x_27, 15); +x_56 = lean_ctor_get(x_27, 16); +x_57 = lean_ctor_get(x_27, 17); +x_58 = lean_ctor_get(x_27, 18); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_27); +x_59 = lean_box(0); +x_60 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_59); +x_61 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_41, x_21, x_60); +x_62 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_62, 0, x_39); +lean_ctor_set(x_62, 1, x_40); +lean_ctor_set(x_62, 2, x_61); +lean_ctor_set(x_62, 3, x_42); +lean_ctor_set(x_62, 4, x_43); +lean_ctor_set(x_62, 5, x_44); +lean_ctor_set(x_62, 6, x_46); +lean_ctor_set(x_62, 7, x_47); +lean_ctor_set(x_62, 8, x_48); +lean_ctor_set(x_62, 9, x_49); +lean_ctor_set(x_62, 10, x_50); +lean_ctor_set(x_62, 11, x_51); +lean_ctor_set(x_62, 12, x_52); +lean_ctor_set(x_62, 13, x_53); +lean_ctor_set(x_62, 14, x_54); +lean_ctor_set(x_62, 15, x_55); +lean_ctor_set(x_62, 16, x_56); +lean_ctor_set(x_62, 17, x_57); +lean_ctor_set(x_62, 18, x_58); +lean_ctor_set_uint8(x_62, sizeof(void*)*19, x_45); +x_63 = lean_st_ref_set(x_3, x_62, x_28); +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_347); - x_378 = lean_box(0); + lean_dec_ref(x_63); + x_65 = lean_box(0); } -x_379 = lean_ctor_get(x_349, 0); -lean_inc(x_379); -x_380 = lean_ctor_get(x_349, 1); -lean_inc(x_380); -x_381 = lean_ctor_get(x_349, 2); -lean_inc(x_381); -x_382 = lean_ctor_get(x_349, 3); -lean_inc(x_382); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_383 = x_349; +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_349); - x_383 = lean_box(0); + x_66 = x_65; } -x_384 = 1; -if (lean_is_scalar(x_383)) { - x_385 = lean_alloc_ctor(1, 4, 1); -} else { - x_385 = x_383; +lean_ctor_set(x_66, 0, x_59); +lean_ctor_set(x_66, 1, x_64); +return x_66; } -lean_ctor_set(x_385, 0, x_379); -lean_ctor_set(x_385, 1, x_380); -lean_ctor_set(x_385, 2, x_381); -lean_ctor_set(x_385, 3, x_382); -lean_ctor_set_uint8(x_385, sizeof(void*)*4, x_384); -if (lean_is_scalar(x_378)) { - x_386 = lean_alloc_ctor(1, 4, 1); -} else { - x_386 = x_378; } -lean_ctor_set(x_386, 0, x_377); -lean_ctor_set(x_386, 1, x_343); -lean_ctor_set(x_386, 2, x_344); -lean_ctor_set(x_386, 3, x_345); -lean_ctor_set_uint8(x_386, sizeof(void*)*4, x_384); -x_387 = 0; -x_388 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_388, 0, x_385); -lean_ctor_set(x_388, 1, x_375); -lean_ctor_set(x_388, 2, x_376); -lean_ctor_set(x_388, 3, x_386); -lean_ctor_set_uint8(x_388, sizeof(void*)*4, x_387); -return x_388; } -else +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_389; -x_389 = lean_ctor_get(x_347, 3); -lean_inc(x_389); -if (lean_obj_tag(x_389) == 0) +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -lean_object* x_390; uint8_t x_391; lean_object* x_392; -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_390 = x_349; -} else { - lean_dec_ref(x_349); - x_390 = lean_box(0); +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_1, x_6, x_7, x_4, x_5); +return x_8; } -x_391 = 1; -if (lean_is_scalar(x_390)) { - x_392 = lean_alloc_ctor(1, 4, 1); -} else { - x_392 = x_390; } -lean_ctor_set(x_392, 0, x_347); -lean_ctor_set(x_392, 1, x_343); -lean_ctor_set(x_392, 2, x_344); -lean_ctor_set(x_392, 3, x_345); -lean_ctor_set_uint8(x_392, sizeof(void*)*4, x_391); -return x_392; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_6; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_393; -x_393 = lean_ctor_get_uint8(x_389, sizeof(void*)*4); -if (x_393 == 0) +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; uint8_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; uint8_t x_406; lean_object* x_407; -x_394 = lean_ctor_get(x_347, 1); -lean_inc(x_394); -x_395 = lean_ctor_get(x_347, 2); -lean_inc(x_395); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_396 = x_347; -} else { - lean_dec_ref(x_347); - x_396 = lean_box(0); +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_1, x_2); +lean_dec(x_2); +return x_3; } -x_397 = lean_ctor_get(x_389, 0); -lean_inc(x_397); -x_398 = lean_ctor_get(x_389, 1); -lean_inc(x_398); -x_399 = lean_ctor_get(x_389, 2); -lean_inc(x_399); -x_400 = lean_ctor_get(x_389, 3); -lean_inc(x_400); -if (lean_is_exclusive(x_389)) { - lean_ctor_release(x_389, 0); - lean_ctor_release(x_389, 1); - lean_ctor_release(x_389, 2); - lean_ctor_release(x_389, 3); - x_401 = x_389; -} else { - lean_dec_ref(x_389); - x_401 = lean_box(0); } -x_402 = 1; -lean_inc(x_349); -if (lean_is_scalar(x_401)) { - x_403 = lean_alloc_ctor(1, 4, 1); -} else { - x_403 = x_401; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_registerParent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } -lean_ctor_set(x_403, 0, x_349); -lean_ctor_set(x_403, 1, x_394); -lean_ctor_set(x_403, 2, x_395); -lean_ctor_set(x_403, 3, x_397); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_404 = x_349; -} else { - lean_dec_ref(x_349); - x_404 = lean_box(0); } -lean_ctor_set_uint8(x_403, sizeof(void*)*4, x_402); -if (lean_is_scalar(x_404)) { - x_405 = lean_alloc_ctor(1, 4, 1); -} else { - x_405 = x_404; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 2); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_14, x_1); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); +return x_11; } -lean_ctor_set(x_405, 0, x_400); -lean_ctor_set(x_405, 1, x_343); -lean_ctor_set(x_405, 2, x_344); -lean_ctor_set(x_405, 3, x_345); -lean_ctor_set_uint8(x_405, sizeof(void*)*4, x_402); -x_406 = 0; -if (lean_is_scalar(x_396)) { - x_407 = lean_alloc_ctor(1, 4, 1); -} else { - x_407 = x_396; +else +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +lean_ctor_set(x_11, 0, x_17); +return x_11; } -lean_ctor_set(x_407, 0, x_403); -lean_ctor_set(x_407, 1, x_398); -lean_ctor_set(x_407, 2, x_399); -lean_ctor_set(x_407, 3, x_405); -lean_ctor_set_uint8(x_407, sizeof(void*)*4, x_406); -return x_407; } else { -lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; uint8_t x_418; lean_object* x_419; -x_408 = lean_ctor_get(x_347, 1); -lean_inc(x_408); -x_409 = lean_ctor_get(x_347, 2); -lean_inc(x_409); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_410 = x_347; -} else { - lean_dec_ref(x_347); - x_410 = lean_box(0); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_11, 0); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_11); +x_20 = lean_ctor_get(x_18, 2); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_20, x_1); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_19); +return x_23; } -x_411 = lean_ctor_get(x_349, 0); -lean_inc(x_411); -x_412 = lean_ctor_get(x_349, 1); -lean_inc(x_412); -x_413 = lean_ctor_get(x_349, 2); -lean_inc(x_413); -x_414 = lean_ctor_get(x_349, 3); -lean_inc(x_414); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_415 = x_349; -} else { - lean_dec_ref(x_349); - x_415 = lean_box(0); +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_19); +return x_25; } -if (lean_is_scalar(x_415)) { - x_416 = lean_alloc_ctor(1, 4, 1); -} else { - x_416 = x_415; } -lean_ctor_set(x_416, 0, x_411); -lean_ctor_set(x_416, 1, x_412); -lean_ctor_set(x_416, 2, x_413); -lean_ctor_set(x_416, 3, x_414); -lean_ctor_set_uint8(x_416, sizeof(void*)*4, x_393); -if (lean_is_scalar(x_410)) { - x_417 = lean_alloc_ctor(1, 4, 1); -} else { - x_417 = x_410; } -lean_ctor_set(x_417, 0, x_416); -lean_ctor_set(x_417, 1, x_408); -lean_ctor_set(x_417, 2, x_409); -lean_ctor_set(x_417, 3, x_389); -lean_ctor_set_uint8(x_417, sizeof(void*)*4, x_348); -x_418 = 1; -x_419 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_419, 0, x_417); -lean_ctor_set(x_419, 1, x_343); -lean_ctor_set(x_419, 2, x_344); -lean_ctor_set(x_419, 3, x_345); -lean_ctor_set_uint8(x_419, sizeof(void*)*4, x_418); -return x_419; } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getParents(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; } } +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_1); +x_5 = lean_nat_dec_lt(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_3); +x_6 = lean_box(0); +return x_6; } +else +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_fget(x_1, x_3); +x_8 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_7, x_2); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_add(x_3, x_9); +lean_dec(x_3); +x_3 = x_10; +goto _start; } else { -uint8_t x_420; lean_object* x_421; -x_420 = 1; -x_421 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_421, 0, x_347); -lean_ctor_set(x_421, 1, x_343); -lean_ctor_set(x_421, 2, x_344); -lean_ctor_set(x_421, 3, x_345); -lean_ctor_set_uint8(x_421, sizeof(void*)*4, x_420); -return x_421; +lean_object* x_12; +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_3); +return x_12; } } -case 1: -{ -uint8_t x_422; lean_object* x_423; -lean_dec(x_344); -lean_dec(x_343); -x_422 = 1; -x_423 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_423, 0, x_342); -lean_ctor_set(x_423, 1, x_2); -lean_ctor_set(x_423, 2, x_3); -lean_ctor_set(x_423, 3, x_345); -lean_ctor_set_uint8(x_423, sizeof(void*)*4, x_422); -return x_423; } -default: +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: { -lean_object* x_424; uint8_t x_425; -x_424 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_345, x_2, x_3); -x_425 = lean_ctor_get_uint8(x_424, sizeof(void*)*4); -if (x_425 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_426; -x_426 = lean_ctor_get(x_424, 0); -lean_inc(x_426); -if (lean_obj_tag(x_426) == 0) +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +switch (lean_obj_tag(x_10)) { +case 0: { -lean_object* x_427; -x_427 = lean_ctor_get(x_424, 3); -lean_inc(x_427); -if (lean_obj_tag(x_427) == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); +lean_dec(x_11); +if (x_12 == 0) { -lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; uint8_t x_432; lean_object* x_433; -x_428 = lean_ctor_get(x_424, 1); -lean_inc(x_428); -x_429 = lean_ctor_get(x_424, 2); -lean_inc(x_429); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_430 = x_424; -} else { - lean_dec_ref(x_424); - x_430 = lean_box(0); -} -if (lean_is_scalar(x_430)) { - x_431 = lean_alloc_ctor(1, 4, 1); -} else { - x_431 = x_430; -} -lean_ctor_set(x_431, 0, x_427); -lean_ctor_set(x_431, 1, x_428); -lean_ctor_set(x_431, 2, x_429); -lean_ctor_set(x_431, 3, x_427); -lean_ctor_set_uint8(x_431, sizeof(void*)*4, x_425); -x_432 = 1; -x_433 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_433, 0, x_342); -lean_ctor_set(x_433, 1, x_343); -lean_ctor_set(x_433, 2, x_344); -lean_ctor_set(x_433, 3, x_431); -lean_ctor_set_uint8(x_433, sizeof(void*)*4, x_432); -return x_433; +lean_dec(x_8); +lean_dec(x_4); +return x_1; } else { -uint8_t x_434; -x_434 = lean_ctor_get_uint8(x_427, sizeof(void*)*4); -if (x_434 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_1); +if (x_13 == 0) { -lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; uint8_t x_443; lean_object* x_444; lean_object* x_445; uint8_t x_446; lean_object* x_447; -x_435 = lean_ctor_get(x_424, 1); -lean_inc(x_435); -x_436 = lean_ctor_get(x_424, 2); -lean_inc(x_436); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_437 = x_424; -} else { - lean_dec_ref(x_424); - x_437 = lean_box(0); -} -x_438 = lean_ctor_get(x_427, 0); -lean_inc(x_438); -x_439 = lean_ctor_get(x_427, 1); -lean_inc(x_439); -x_440 = lean_ctor_get(x_427, 2); -lean_inc(x_440); -x_441 = lean_ctor_get(x_427, 3); -lean_inc(x_441); -if (lean_is_exclusive(x_427)) { - lean_ctor_release(x_427, 0); - lean_ctor_release(x_427, 1); - lean_ctor_release(x_427, 2); - lean_ctor_release(x_427, 3); - x_442 = x_427; -} else { - lean_dec_ref(x_427); - x_442 = lean_box(0); -} -x_443 = 1; -if (lean_is_scalar(x_442)) { - x_444 = lean_alloc_ctor(1, 4, 1); -} else { - x_444 = x_442; -} -lean_ctor_set(x_444, 0, x_342); -lean_ctor_set(x_444, 1, x_343); -lean_ctor_set(x_444, 2, x_344); -lean_ctor_set(x_444, 3, x_426); -lean_ctor_set_uint8(x_444, sizeof(void*)*4, x_443); -if (lean_is_scalar(x_437)) { - x_445 = lean_alloc_ctor(1, 4, 1); -} else { - x_445 = x_437; -} -lean_ctor_set(x_445, 0, x_438); -lean_ctor_set(x_445, 1, x_439); -lean_ctor_set(x_445, 2, x_440); -lean_ctor_set(x_445, 3, x_441); -lean_ctor_set_uint8(x_445, sizeof(void*)*4, x_443); -x_446 = 0; -x_447 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_447, 0, x_444); -lean_ctor_set(x_447, 1, x_435); -lean_ctor_set(x_447, 2, x_436); -lean_ctor_set(x_447, 3, x_445); -lean_ctor_set_uint8(x_447, sizeof(void*)*4, x_446); -return x_447; +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_1, 0); +lean_dec(x_14); +x_15 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_15); +return x_1; } else { -lean_object* x_448; uint8_t x_449; lean_object* x_450; -if (lean_is_exclusive(x_427)) { - lean_ctor_release(x_427, 0); - lean_ctor_release(x_427, 1); - lean_ctor_release(x_427, 2); - lean_ctor_release(x_427, 3); - x_448 = x_427; -} else { - lean_dec_ref(x_427); - x_448 = lean_box(0); -} -x_449 = 1; -if (lean_is_scalar(x_448)) { - x_450 = lean_alloc_ctor(1, 4, 1); -} else { - x_450 = x_448; +lean_object* x_16; lean_object* x_17; +lean_dec(x_1); +x_16 = lean_array_set(x_4, x_8, x_9); +lean_dec(x_8); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; } -lean_ctor_set(x_450, 0, x_342); -lean_ctor_set(x_450, 1, x_343); -lean_ctor_set(x_450, 2, x_344); -lean_ctor_set(x_450, 3, x_424); -lean_ctor_set_uint8(x_450, sizeof(void*)*4, x_449); -return x_450; } } +case 1: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_1); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_ctor_get(x_1, 0); +lean_dec(x_19); +x_20 = !lean_is_exclusive(x_10); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_ctor_get(x_10, 0); +x_22 = lean_array_set(x_4, x_8, x_9); +x_23 = lean_usize_shift_right(x_2, x_5); +x_24 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_21, x_23, x_3); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; +lean_ctor_set(x_10, 0, x_24); +x_26 = lean_array_set(x_22, x_8, x_10); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_26); +return x_1; } else { -uint8_t x_451; -x_451 = lean_ctor_get_uint8(x_426, sizeof(void*)*4); -if (x_451 == 0) +lean_object* x_27; uint8_t x_28; +lean_dec(x_24); +lean_free_object(x_10); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; uint8_t x_461; lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; -x_452 = lean_ctor_get(x_424, 1); -lean_inc(x_452); -x_453 = lean_ctor_get(x_424, 2); -lean_inc(x_453); -x_454 = lean_ctor_get(x_424, 3); -lean_inc(x_454); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_455 = x_424; -} else { - lean_dec_ref(x_424); - x_455 = lean_box(0); -} -x_456 = lean_ctor_get(x_426, 0); -lean_inc(x_456); -x_457 = lean_ctor_get(x_426, 1); -lean_inc(x_457); -x_458 = lean_ctor_get(x_426, 2); -lean_inc(x_458); -x_459 = lean_ctor_get(x_426, 3); -lean_inc(x_459); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_460 = x_426; -} else { - lean_dec_ref(x_426); - x_460 = lean_box(0); +lean_object* x_29; +x_29 = lean_array_set(x_22, x_8, x_27); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_29); +return x_1; } -x_461 = 1; -if (lean_is_scalar(x_460)) { - x_462 = lean_alloc_ctor(1, 4, 1); -} else { - x_462 = x_460; +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_27, 0); +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_27); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_array_set(x_22, x_8, x_32); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_33); +return x_1; } -lean_ctor_set(x_462, 0, x_342); -lean_ctor_set(x_462, 1, x_343); -lean_ctor_set(x_462, 2, x_344); -lean_ctor_set(x_462, 3, x_456); -lean_ctor_set_uint8(x_462, sizeof(void*)*4, x_461); -if (lean_is_scalar(x_455)) { - x_463 = lean_alloc_ctor(1, 4, 1); -} else { - x_463 = x_455; } -lean_ctor_set(x_463, 0, x_459); -lean_ctor_set(x_463, 1, x_452); -lean_ctor_set(x_463, 2, x_453); -lean_ctor_set(x_463, 3, x_454); -lean_ctor_set_uint8(x_463, sizeof(void*)*4, x_461); -x_464 = 0; -x_465 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_465, 0, x_462); -lean_ctor_set(x_465, 1, x_457); -lean_ctor_set(x_465, 2, x_458); -lean_ctor_set(x_465, 3, x_463); -lean_ctor_set_uint8(x_465, sizeof(void*)*4, x_464); -return x_465; } else { -lean_object* x_466; -x_466 = lean_ctor_get(x_424, 3); -lean_inc(x_466); -if (lean_obj_tag(x_466) == 0) +lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; +x_34 = lean_ctor_get(x_10, 0); +lean_inc(x_34); +lean_dec(x_10); +x_35 = lean_array_set(x_4, x_8, x_9); +x_36 = lean_usize_shift_right(x_2, x_5); +x_37 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_34, x_36, x_3); +lean_inc(x_37); +x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_467; uint8_t x_468; lean_object* x_469; -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_467 = x_426; -} else { - lean_dec_ref(x_426); - x_467 = lean_box(0); -} -x_468 = 1; -if (lean_is_scalar(x_467)) { - x_469 = lean_alloc_ctor(1, 4, 1); -} else { - x_469 = x_467; -} -lean_ctor_set(x_469, 0, x_342); -lean_ctor_set(x_469, 1, x_343); -lean_ctor_set(x_469, 2, x_344); -lean_ctor_set(x_469, 3, x_424); -lean_ctor_set_uint8(x_469, sizeof(void*)*4, x_468); -return x_469; +lean_object* x_39; lean_object* x_40; +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_37); +x_40 = lean_array_set(x_35, x_8, x_39); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_40); +return x_1; } else { -uint8_t x_470; -x_470 = lean_ctor_get_uint8(x_466, sizeof(void*)*4); -if (x_470 == 0) -{ -lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; uint8_t x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; uint8_t x_483; lean_object* x_484; -x_471 = lean_ctor_get(x_424, 1); -lean_inc(x_471); -x_472 = lean_ctor_get(x_424, 2); -lean_inc(x_472); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_473 = x_424; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_37); +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(x_38); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_44 = x_41; } else { - lean_dec_ref(x_424); - x_473 = lean_box(0); + lean_dec_ref(x_41); + x_44 = lean_box(0); } -x_474 = lean_ctor_get(x_466, 0); -lean_inc(x_474); -x_475 = lean_ctor_get(x_466, 1); -lean_inc(x_475); -x_476 = lean_ctor_get(x_466, 2); -lean_inc(x_476); -x_477 = lean_ctor_get(x_466, 3); -lean_inc(x_477); -if (lean_is_exclusive(x_466)) { - lean_ctor_release(x_466, 0); - lean_ctor_release(x_466, 1); - lean_ctor_release(x_466, 2); - lean_ctor_release(x_466, 3); - x_478 = x_466; +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_466); - x_478 = lean_box(0); + x_45 = x_44; } -x_479 = 1; -lean_inc(x_426); -if (lean_is_scalar(x_478)) { - x_480 = lean_alloc_ctor(1, 4, 1); -} else { - x_480 = x_478; +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +x_46 = lean_array_set(x_35, x_8, x_45); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_46); +return x_1; } -lean_ctor_set(x_480, 0, x_342); -lean_ctor_set(x_480, 1, x_343); -lean_ctor_set(x_480, 2, x_344); -lean_ctor_set(x_480, 3, x_426); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_481 = x_426; -} else { - lean_dec_ref(x_426); - x_481 = lean_box(0); } -lean_ctor_set_uint8(x_480, sizeof(void*)*4, x_479); -if (lean_is_scalar(x_481)) { - x_482 = lean_alloc_ctor(1, 4, 1); -} else { - x_482 = x_481; } -lean_ctor_set(x_482, 0, x_474); -lean_ctor_set(x_482, 1, x_475); -lean_ctor_set(x_482, 2, x_476); -lean_ctor_set(x_482, 3, x_477); -lean_ctor_set_uint8(x_482, sizeof(void*)*4, x_479); -x_483 = 0; -if (lean_is_scalar(x_473)) { - x_484 = lean_alloc_ctor(1, 4, 1); +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_1); +x_47 = lean_ctor_get(x_10, 0); +lean_inc(x_47); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_48 = x_10; } else { - x_484 = x_473; + lean_dec_ref(x_10); + x_48 = lean_box(0); } -lean_ctor_set(x_484, 0, x_480); -lean_ctor_set(x_484, 1, x_471); -lean_ctor_set(x_484, 2, x_472); -lean_ctor_set(x_484, 3, x_482); -lean_ctor_set_uint8(x_484, sizeof(void*)*4, x_483); -return x_484; +x_49 = lean_array_set(x_4, x_8, x_9); +x_50 = lean_usize_shift_right(x_2, x_5); +x_51 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_47, x_50, x_3); +lean_inc(x_51); +x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +if (lean_is_scalar(x_48)) { + x_53 = lean_alloc_ctor(1, 1, 0); +} else { + x_53 = x_48; +} +lean_ctor_set(x_53, 0, x_51); +x_54 = lean_array_set(x_49, x_8, x_53); +lean_dec(x_8); +x_55 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_55, 0, x_54); +return x_55; } else { -lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; -x_485 = lean_ctor_get(x_424, 1); -lean_inc(x_485); -x_486 = lean_ctor_get(x_424, 2); -lean_inc(x_486); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_487 = x_424; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_51); +lean_dec(x_48); +x_56 = lean_ctor_get(x_52, 0); +lean_inc(x_56); +lean_dec(x_52); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; } else { - lean_dec_ref(x_424); - x_487 = lean_box(0); + lean_dec_ref(x_56); + x_59 = lean_box(0); } -x_488 = lean_ctor_get(x_426, 0); -lean_inc(x_488); -x_489 = lean_ctor_get(x_426, 1); -lean_inc(x_489); -x_490 = lean_ctor_get(x_426, 2); -lean_inc(x_490); -x_491 = lean_ctor_get(x_426, 3); -lean_inc(x_491); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_492 = x_426; +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_426); - x_492 = lean_box(0); + x_60 = x_59; } -if (lean_is_scalar(x_492)) { - x_493 = lean_alloc_ctor(1, 4, 1); -} else { - x_493 = x_492; +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_58); +x_61 = lean_array_set(x_49, x_8, x_60); +lean_dec(x_8); +x_62 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_62, 0, x_61); +return x_62; } -lean_ctor_set(x_493, 0, x_488); -lean_ctor_set(x_493, 1, x_489); -lean_ctor_set(x_493, 2, x_490); -lean_ctor_set(x_493, 3, x_491); -lean_ctor_set_uint8(x_493, sizeof(void*)*4, x_470); -if (lean_is_scalar(x_487)) { - x_494 = lean_alloc_ctor(1, 4, 1); -} else { - x_494 = x_487; } -lean_ctor_set(x_494, 0, x_493); -lean_ctor_set(x_494, 1, x_485); -lean_ctor_set(x_494, 2, x_486); -lean_ctor_set(x_494, 3, x_466); -lean_ctor_set_uint8(x_494, sizeof(void*)*4, x_425); -x_495 = 1; -x_496 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_496, 0, x_342); -lean_ctor_set(x_496, 1, x_343); -lean_ctor_set(x_496, 2, x_344); -lean_ctor_set(x_496, 3, x_494); -lean_ctor_set_uint8(x_496, sizeof(void*)*4, x_495); -return x_496; } +default: +{ +lean_dec(x_8); +lean_dec(x_4); +return x_1; } } } +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_1, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 1); +lean_inc(x_64); +x_65 = lean_unsigned_to_nat(0u); +x_66 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_63, x_3, x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_dec(x_64); +lean_dec(x_63); +return x_1; } else { -uint8_t x_497; lean_object* x_498; -x_497 = 1; -x_498 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_498, 0, x_342); -lean_ctor_set(x_498, 1, x_343); -lean_ctor_set(x_498, 2, x_344); -lean_ctor_set(x_498, 3, x_424); -lean_ctor_set_uint8(x_498, sizeof(void*)*4, x_497); -return x_498; +uint8_t x_67; +x_67 = !lean_is_exclusive(x_1); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_68 = lean_ctor_get(x_1, 1); +lean_dec(x_68); +x_69 = lean_ctor_get(x_1, 0); +lean_dec(x_69); +x_70 = lean_ctor_get(x_66, 0); +lean_inc(x_70); +lean_dec(x_66); +lean_inc(x_70); +x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); +x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); +lean_ctor_set(x_1, 1, x_72); +lean_ctor_set(x_1, 0, x_71); +return x_1; } +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_1); +x_73 = lean_ctor_get(x_66, 0); +lean_inc(x_73); +lean_dec(x_66); +lean_inc(x_73); +x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); +x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; size_t x_4; lean_object* x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_4; -x_4 = l_Lean_RBNode_isRed___rarg(x_1); -if (x_4 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_11 = l_Lean_Meta_Grind_getParents(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_st_ref_take(x_2, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) { -lean_object* x_5; -x_5 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); -return x_5; +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_ctor_get(x_15, 2); +x_19 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_18, x_1); +lean_ctor_set(x_15, 2, x_19); +x_20 = lean_st_ref_set(x_2, x_15, x_16); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +lean_ctor_set(x_20, 0, x_12); +return x_20; } else { -lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_RBNode_ins___at_Lean_Meta_Grind_registerParent___spec__2(x_1, x_2, x_3); -x_7 = l_Lean_RBNode_setBlack___rarg(x_6); -return x_7; +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_12); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_25 = lean_ctor_get(x_15, 0); +x_26 = lean_ctor_get(x_15, 1); +x_27 = lean_ctor_get(x_15, 2); +x_28 = lean_ctor_get(x_15, 3); +x_29 = lean_ctor_get(x_15, 4); +x_30 = lean_ctor_get(x_15, 5); +x_31 = lean_ctor_get_uint8(x_15, sizeof(void*)*19); +x_32 = lean_ctor_get(x_15, 6); +x_33 = lean_ctor_get(x_15, 7); +x_34 = lean_ctor_get(x_15, 8); +x_35 = lean_ctor_get(x_15, 9); +x_36 = lean_ctor_get(x_15, 10); +x_37 = lean_ctor_get(x_15, 11); +x_38 = lean_ctor_get(x_15, 12); +x_39 = lean_ctor_get(x_15, 13); +x_40 = lean_ctor_get(x_15, 14); +x_41 = lean_ctor_get(x_15, 15); +x_42 = lean_ctor_get(x_15, 16); +x_43 = lean_ctor_get(x_15, 17); +x_44 = lean_ctor_get(x_15, 18); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_15); +x_45 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_27, x_1); +x_46 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_46, 0, x_25); +lean_ctor_set(x_46, 1, x_26); +lean_ctor_set(x_46, 2, x_45); +lean_ctor_set(x_46, 3, x_28); +lean_ctor_set(x_46, 4, x_29); +lean_ctor_set(x_46, 5, x_30); +lean_ctor_set(x_46, 6, x_32); +lean_ctor_set(x_46, 7, x_33); +lean_ctor_set(x_46, 8, x_34); +lean_ctor_set(x_46, 9, x_35); +lean_ctor_set(x_46, 10, x_36); +lean_ctor_set(x_46, 11, x_37); +lean_ctor_set(x_46, 12, x_38); +lean_ctor_set(x_46, 13, x_39); +lean_ctor_set(x_46, 14, x_40); +lean_ctor_set(x_46, 15, x_41); +lean_ctor_set(x_46, 16, x_42); +lean_ctor_set(x_46, 17, x_43); +lean_ctor_set(x_46, 18, x_44); +lean_ctor_set_uint8(x_46, sizeof(void*)*19, x_31); +x_47 = lean_st_ref_set(x_2, x_46, x_16); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); +} +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(0, 2, 0); +} else { + x_50 = x_49; } +lean_ctor_set(x_50, 0, x_12); +lean_ctor_set(x_50, 1, x_48); +return x_50; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_5); -return x_6; } -else +LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_9); -x_12 = lean_uint64_to_usize(x_11); -x_13 = 1; -x_14 = lean_usize_sub(x_1, x_13); -x_15 = 5; -x_16 = lean_usize_mul(x_15, x_14); -x_17 = lean_usize_shift_right(x_12, x_16); -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_5, x_18); -lean_dec(x_5); -x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_6, x_17, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_19; -x_6 = x_20; -goto _start; -} +lean_object* x_4; +x_4 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -uint8_t x_9; +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_1, 1); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 0); -lean_dec(x_11); -x_12 = lean_array_push(x_5, x_3); -x_13 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_12); -return x_1; +x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_1); -x_14 = lean_array_push(x_5, x_3); -x_15 = lean_array_push(x_6, x_4); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_object* x_3; +x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; } } -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_5, x_2); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_19; lean_object* x_20; +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getParentsAndReset(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_2 = x_20; -goto _start; +lean_dec(x_1); +return x_11; } -else +} +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -if (x_22 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 1); -lean_dec(x_23); -x_24 = lean_ctor_get(x_1, 0); -lean_dec(x_24); -x_25 = lean_array_fset(x_5, x_2, x_3); -x_26 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -lean_ctor_set(x_1, 1, x_26); -lean_ctor_set(x_1, 0, x_25); -return x_1; +lean_object* x_12; lean_object* x_13; +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 3); +lean_inc(x_16); lean_dec(x_1); -x_27 = lean_array_fset(x_5, x_2, x_3); -x_28 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} -} +x_17 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_box(0); +x_22 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_20, x_15, x_21); +x_1 = x_16; +x_2 = x_22; +x_11 = x_19; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = 1; -x_9 = 5; -x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_14 = lean_nat_dec_lt(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) -{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_59; +x_12 = lean_st_ref_get(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_4); -return x_1; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_19); -if (x_21 == 0) +x_15 = lean_ctor_get(x_13, 2); +lean_inc(x_15); +lean_dec(x_13); +x_59 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_15, x_2); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_free_object(x_15); -x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_array_fset(x_17, x_12, x_23); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_24); -return x_1; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = lean_box(0); +x_61 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_ctor_get(x_62, 0); +lean_inc(x_64); +lean_dec(x_62); +x_16 = x_64; +x_17 = x_63; +goto block_58; } else { -lean_object* x_25; -lean_dec(x_20); -lean_dec(x_19); -lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_4); -x_25 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_25); -return x_1; -} +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_65 = lean_ctor_get(x_59, 0); +lean_inc(x_65); +lean_dec(x_59); +x_66 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_65, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_ctor_get(x_67, 0); +lean_inc(x_69); +lean_dec(x_67); +x_16 = x_69; +x_17 = x_68; +goto block_58; } -else +block_58: { -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_26); -if (x_28 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_st_ref_take(x_3, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_array_fset(x_17, x_12, x_30); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); -return x_1; -} -else +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_19, 2); +x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_22, x_2, x_16); +lean_ctor_set(x_19, 2, x_23); +x_24 = lean_st_ref_set(x_3, x_19, x_20); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -lean_object* x_32; lean_object* x_33; -lean_dec(x_27); +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); lean_dec(x_26); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_4); -lean_ctor_set(x_32, 1, x_5); -x_33 = lean_array_fset(x_17, x_12, x_32); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_33); -return x_1; -} -} +x_27 = lean_box(0); +lean_ctor_set(x_24, 0, x_27); +return x_24; } -case 1: -{ -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) +else { -lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -x_35 = lean_ctor_get(x_15, 0); -x_36 = lean_usize_shift_right(x_2, x_9); -x_37 = lean_usize_add(x_3, x_8); -x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_35, x_36, x_37, x_4, x_5); -lean_ctor_set(x_15, 0, x_38); -x_39 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); -return x_1; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_24, 1); +lean_inc(x_28); +lean_dec(x_24); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; +} } else { -lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_15, 0); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_31 = lean_ctor_get(x_19, 0); +x_32 = lean_ctor_get(x_19, 1); +x_33 = lean_ctor_get(x_19, 2); +x_34 = lean_ctor_get(x_19, 3); +x_35 = lean_ctor_get(x_19, 4); +x_36 = lean_ctor_get(x_19, 5); +x_37 = lean_ctor_get_uint8(x_19, sizeof(void*)*19); +x_38 = lean_ctor_get(x_19, 6); +x_39 = lean_ctor_get(x_19, 7); +x_40 = lean_ctor_get(x_19, 8); +x_41 = lean_ctor_get(x_19, 9); +x_42 = lean_ctor_get(x_19, 10); +x_43 = lean_ctor_get(x_19, 11); +x_44 = lean_ctor_get(x_19, 12); +x_45 = lean_ctor_get(x_19, 13); +x_46 = lean_ctor_get(x_19, 14); +x_47 = lean_ctor_get(x_19, 15); +x_48 = lean_ctor_get(x_19, 16); +x_49 = lean_ctor_get(x_19, 17); +x_50 = lean_ctor_get(x_19, 18); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); lean_inc(x_40); -lean_dec(x_15); -x_41 = lean_usize_shift_right(x_2, x_9); -x_42 = lean_usize_add(x_3, x_8); -x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_40, x_41, x_42, x_4, x_5); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_array_fset(x_17, x_12, x_44); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); -return x_1; +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_19); +x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_33, x_2, x_16); +x_52 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_52, 0, x_31); +lean_ctor_set(x_52, 1, x_32); +lean_ctor_set(x_52, 2, x_51); +lean_ctor_set(x_52, 3, x_34); +lean_ctor_set(x_52, 4, x_35); +lean_ctor_set(x_52, 5, x_36); +lean_ctor_set(x_52, 6, x_38); +lean_ctor_set(x_52, 7, x_39); +lean_ctor_set(x_52, 8, x_40); +lean_ctor_set(x_52, 9, x_41); +lean_ctor_set(x_52, 10, x_42); +lean_ctor_set(x_52, 11, x_43); +lean_ctor_set(x_52, 12, x_44); +lean_ctor_set(x_52, 13, x_45); +lean_ctor_set(x_52, 14, x_46); +lean_ctor_set(x_52, 15, x_47); +lean_ctor_set(x_52, 16, x_48); +lean_ctor_set(x_52, 17, x_49); +lean_ctor_set(x_52, 18, x_50); +lean_ctor_set_uint8(x_52, sizeof(void*)*19, x_37); +x_53 = lean_st_ref_set(x_3, x_52, x_20); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_55 = x_53; +} else { + lean_dec_ref(x_53); + x_55 = lean_box(0); } +x_56 = lean_box(0); +if (lean_is_scalar(x_55)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_55; } -default: -{ -lean_object* x_46; lean_object* x_47; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_4); -lean_ctor_set(x_46, 1, x_5); -x_47 = lean_array_fset(x_17, x_12, x_46); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_47); -return x_1; +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +return x_57; } } } } -else -{ -lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_1, 0); -lean_inc(x_48); -lean_dec(x_1); -x_49 = 1; -x_50 = 5; -x_51 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_52 = lean_usize_land(x_2, x_51); -x_53 = lean_usize_to_nat(x_52); -x_54 = lean_array_get_size(x_48); -x_55 = lean_nat_dec_lt(x_53, x_54); -lean_dec(x_54); -if (x_55 == 0) +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_56; -lean_dec(x_53); +lean_object* x_12; +x_12 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_48); -return x_56; -} -else -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_array_fget(x_48, x_53); -x_58 = lean_box(0); -x_59 = lean_array_fset(x_48, x_53, x_58); -switch (lean_obj_tag(x_57)) { -case 0: -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_57, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_62 = x_57; -} else { - lean_dec_ref(x_57); - x_62 = lean_box(0); +lean_dec(x_3); +return x_12; } -x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_array_fset(x_59, x_53, x_65); -lean_dec(x_53); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_61); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_68 = lean_alloc_ctor(0, 2, 0); -} else { - x_68 = x_62; -} -lean_ctor_set(x_68, 0, x_4); -lean_ctor_set(x_68, 1, x_5); -x_69 = lean_array_fset(x_59, x_53, x_68); -lean_dec(x_53); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_copyParentsTo(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; } } -case 1: +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg(lean_object* x_1) { +_start: { -lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_71 = lean_ctor_get(x_57, 0); -lean_inc(x_71); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - x_72 = x_57; -} else { - lean_dec_ref(x_57); - x_72 = lean_box(0); -} -x_73 = lean_usize_shift_right(x_2, x_50); -x_74 = lean_usize_add(x_3, x_49); -x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_71, x_73, x_74, x_4, x_5); -if (lean_is_scalar(x_72)) { - x_76 = lean_alloc_ctor(1, 1, 0); -} else { - x_76 = x_72; +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +return x_2; } -lean_ctor_set(x_76, 0, x_75); -x_77 = lean_array_fset(x_59, x_53, x_76); -lean_dec(x_53); -x_78 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_78, 0, x_77); -return x_78; } -default: +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_4); -lean_ctor_set(x_79, 1, x_5); -x_80 = lean_array_fset(x_59, x_53, x_79); -lean_dec(x_53); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; -} +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed), 1, 0); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_setENode_unsafe__1___rarg(x_1); +lean_dec(x_1); +return x_2; } } -else -{ -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; -x_83 = lean_unsigned_to_nat(0u); -x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(x_1, x_83, x_4, x_5); -x_85 = 7; -x_86 = lean_usize_dec_le(x_85, x_3); -if (x_86 == 0) +lean_object* x_3; +x_3 = l_Lean_Meta_Grind_setENode_unsafe__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_dec_lt(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_84, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 1); -lean_inc(x_91); -lean_dec(x_84); -x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(x_3, x_90, x_91, lean_box(0), x_83, x_92); -lean_dec(x_91); -lean_dec(x_90); -return x_93; +lean_dec(x_5); +return x_6; } else { -return x_84; +lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_9); +x_12 = lean_uint64_to_usize(x_11); +x_13 = 1; +x_14 = lean_usize_sub(x_1, x_13); +x_15 = 5; +x_16 = lean_usize_mul(x_15, x_14); +x_17 = lean_usize_shift_right(x_12, x_16); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_5, x_18); +lean_dec(x_5); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_6, x_17, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_19; +x_6 = x_20; +goto _start; } } -else -{ -return x_84; } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; -x_94 = lean_ctor_get(x_1, 0); -x_95 = lean_ctor_get(x_1, 1); -lean_inc(x_95); -lean_inc(x_94); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_1); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -x_97 = lean_unsigned_to_nat(0u); -x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_registerParent___spec__6(x_96, x_97, x_4, x_5); -x_99 = 7; -x_100 = lean_usize_dec_le(x_99, x_3); -if (x_100 == 0) +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +else { -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); -x_102 = lean_unsigned_to_nat(4u); -x_103 = lean_nat_dec_lt(x_101, x_102); -lean_dec(x_101); -if (x_103 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_98, 1); -lean_inc(x_105); -lean_dec(x_98); -x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; +lean_object* x_19; lean_object* x_20; +lean_dec(x_6); +lean_dec(x_5); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_2); +x_2 = x_20; +goto _start; } else { -return x_98; -} +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_1, 0); +lean_dec(x_24); +x_25 = lean_array_fset(x_5, x_2, x_3); +x_26 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); +return x_1; } else { -return x_98; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_1, x_5, x_6, x_2, x_3); -return x_7; -} +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_19); +if (x_21 == 0) { -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); -if (x_7 == 0) +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_array_fset(x_17, x_12, x_23); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_24); +return x_1; +} +else { -lean_object* x_8; -lean_dec(x_4); -x_8 = lean_box(0); -return x_8; +lean_object* x_25; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_25 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_25); +return x_1; +} } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); -lean_dec(x_9); -if (x_10 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_26); +if (x_28 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_12; -goto _start; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = lean_array_fset(x_17, x_12, x_30); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_31); +return x_1; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_array_fget(x_2, x_4); -lean_dec(x_4); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; -} +lean_object* x_32; lean_object* x_33; +lean_dec(x_27); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_5); +x_33 = lean_array_fset(x_17, x_12, x_32); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_33); +return x_1; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +case 1: { -lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_1, 0); -x_6 = 5; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_8 = lean_usize_land(x_2, x_7); -x_9 = lean_usize_to_nat(x_8); -x_10 = lean_box(2); -x_11 = lean_array_get(x_10, x_5, x_9); -lean_dec(x_9); -lean_dec(x_5); -switch (lean_obj_tag(x_11)) { -case 0: +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_12); +lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_usize_shift_right(x_2, x_9); +x_37 = lean_usize_add(x_3, x_8); +x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_35, x_36, x_37, x_4, x_5); +lean_ctor_set(x_15, 0, x_38); +x_39 = lean_array_fset(x_17, x_12, x_15); lean_dec(x_12); -if (x_14 == 0) -{ -lean_object* x_15; -lean_dec(x_13); -lean_free_object(x_1); -x_15 = lean_box(0); -return x_15; +lean_ctor_set(x_1, 0, x_39); +return x_1; } else { -lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_13); +lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_15, 0); +lean_inc(x_40); +lean_dec(x_15); +x_41 = lean_usize_shift_right(x_2, x_9); +x_42 = lean_usize_add(x_3, x_8); +x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_40, x_41, x_42, x_4, x_5); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_array_fset(x_17, x_12, x_44); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_45); return x_1; } } -case 1: -{ -lean_object* x_16; size_t x_17; -lean_free_object(x_1); -x_16 = lean_ctor_get(x_11, 0); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_usize_shift_right(x_2, x_6); -x_1 = x_16; -x_2 = x_17; -goto _start; -} default: { -lean_object* x_19; -lean_free_object(x_1); -x_19 = lean_box(0); -return x_19; +lean_object* x_46; lean_object* x_47; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_4); +lean_ctor_set(x_46, 1, x_5); +x_47 = lean_array_fset(x_17, x_12, x_46); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_47); +return x_1; +} } } } else { -lean_object* x_20; size_t x_21; size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); +lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); lean_dec(x_1); -x_21 = 5; -x_22 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_23 = lean_usize_land(x_2, x_22); -x_24 = lean_usize_to_nat(x_23); -x_25 = lean_box(2); -x_26 = lean_array_get(x_25, x_20, x_24); -lean_dec(x_24); -lean_dec(x_20); -switch (lean_obj_tag(x_26)) { +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_52 = lean_usize_land(x_2, x_51); +x_53 = lean_usize_to_nat(x_52); +x_54 = lean_array_get_size(x_48); +x_55 = lean_nat_dec_lt(x_53, x_54); +lean_dec(x_54); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_53); +lean_dec(x_5); +lean_dec(x_4); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_array_fget(x_48, x_53); +x_58 = lean_box(0); +x_59 = lean_array_fset(x_48, x_53, x_58); +switch (lean_obj_tag(x_57)) { case 0: { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_27); -lean_dec(x_27); -if (x_29 == 0) +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_62 = x_57; +} else { + lean_dec_ref(x_57); + x_62 = lean_box(0); +} +x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); +if (x_63 == 0) { -lean_object* x_30; -lean_dec(x_28); -x_30 = lean_box(0); -return x_30; +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_66 = lean_array_fset(x_59, x_53, x_65); +lean_dec(x_53); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } else { -lean_object* x_31; -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_28); -return x_31; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_61); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_62; +} +lean_ctor_set(x_68, 0, x_4); +lean_ctor_set(x_68, 1, x_5); +x_69 = lean_array_fset(x_59, x_53, x_68); +lean_dec(x_53); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; } } case 1: { -lean_object* x_32; size_t x_33; -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_usize_shift_right(x_2, x_21); -x_1 = x_32; -x_2 = x_33; -goto _start; +lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_71 = lean_ctor_get(x_57, 0); +lean_inc(x_71); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_72 = x_57; +} else { + lean_dec_ref(x_57); + x_72 = lean_box(0); +} +x_73 = lean_usize_shift_right(x_2, x_50); +x_74 = lean_usize_add(x_3, x_49); +x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_71, x_73, x_74, x_4, x_5); +if (lean_is_scalar(x_72)) { + x_76 = lean_alloc_ctor(1, 1, 0); +} else { + x_76 = x_72; +} +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_array_fset(x_59, x_53, x_76); +lean_dec(x_53); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; } default: { -lean_object* x_35; -x_35 = lean_box(0); -return x_35; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_4); +lean_ctor_set(x_79, 1, x_5); +x_80 = lean_array_fset(x_59, x_53, x_79); +lean_dec(x_53); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} } } } } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_1, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_1, 1); -lean_inc(x_37); -lean_dec(x_1); -x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(x_36, x_37, lean_box(0), x_38, x_3); -lean_dec(x_37); -lean_dec(x_36); -return x_39; +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; +x_83 = lean_unsigned_to_nat(0u); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(x_1, x_83, x_4, x_5); +x_85 = 7; +x_86 = lean_usize_dec_le(x_85, x_3); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_dec_lt(x_87, x_88); +lean_dec(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_84, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 1); +lean_inc(x_91); +lean_dec(x_84); +x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } +else +{ +return x_84; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(x_1, x_4, x_2); -return x_5; +return x_84; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_Meta_Grind_getRoot_x3f(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -if (lean_obj_tag(x_13) == 0) +else { -uint8_t x_14; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); +lean_inc(x_94); lean_dec(x_1); -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_unsigned_to_nat(0u); +x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(x_96, x_97, x_4, x_5); +x_99 = 7; +x_100 = lean_usize_dec_le(x_99, x_3); +if (x_100 == 0) { -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_12, 0); -lean_dec(x_15); -x_16 = lean_box(0); -lean_ctor_set(x_12, 0, x_16); -return x_12; +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_dec_lt(x_101, x_102); +lean_dec(x_101); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; +return x_98; } } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_63; lean_object* x_64; -x_20 = lean_ctor_get(x_12, 1); -lean_inc(x_20); -lean_dec(x_12); -x_21 = lean_ctor_get(x_13, 0); -lean_inc(x_21); -lean_dec(x_13); -x_22 = lean_st_ref_get(x_3, x_20); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_23, 2); -lean_inc(x_25); -lean_dec(x_23); -x_63 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_25, x_21); -x_64 = lean_st_ref_take(x_3, x_24); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = lean_box(0); -x_26 = x_67; -x_27 = x_65; -x_28 = x_66; -goto block_62; +return x_98; } -else +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_63, 0); -lean_inc(x_68); -lean_dec(x_63); -x_69 = lean_ctor_get(x_64, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_64, 1); -lean_inc(x_70); -lean_dec(x_64); -x_26 = x_68; -x_27 = x_69; -x_28 = x_70; -goto block_62; +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_5, x_6, x_2, x_3); +return x_7; } -block_62: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_27); -if (x_29 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_30 = lean_ctor_get(x_27, 2); -x_31 = lean_box(0); -x_32 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_31); -x_33 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_30, x_21, x_32); -lean_ctor_set(x_27, 2, x_33); -x_34 = lean_st_ref_set(x_3, x_27, x_28); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_13, 1); +x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_16, x_1, x_2); +lean_ctor_set(x_13, 1, x_17); +x_18 = lean_st_ref_set(x_3, x_13, x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -lean_ctor_set(x_34, 0, x_31); -return x_34; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; } else { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); -lean_dec(x_34); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_31); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; } } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_39 = lean_ctor_get(x_27, 0); -x_40 = lean_ctor_get(x_27, 1); -x_41 = lean_ctor_get(x_27, 2); -x_42 = lean_ctor_get(x_27, 3); -x_43 = lean_ctor_get(x_27, 4); -x_44 = lean_ctor_get(x_27, 5); -x_45 = lean_ctor_get_uint8(x_27, sizeof(void*)*14); -x_46 = lean_ctor_get(x_27, 6); -x_47 = lean_ctor_get(x_27, 7); -x_48 = lean_ctor_get(x_27, 8); -x_49 = lean_ctor_get(x_27, 9); -x_50 = lean_ctor_get(x_27, 10); -x_51 = lean_ctor_get(x_27, 11); -x_52 = lean_ctor_get(x_27, 12); -x_53 = lean_ctor_get(x_27, 13); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_25 = lean_ctor_get(x_13, 0); +x_26 = lean_ctor_get(x_13, 1); +x_27 = lean_ctor_get(x_13, 2); +x_28 = lean_ctor_get(x_13, 3); +x_29 = lean_ctor_get(x_13, 4); +x_30 = lean_ctor_get(x_13, 5); +x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*19); +x_32 = lean_ctor_get(x_13, 6); +x_33 = lean_ctor_get(x_13, 7); +x_34 = lean_ctor_get(x_13, 8); +x_35 = lean_ctor_get(x_13, 9); +x_36 = lean_ctor_get(x_13, 10); +x_37 = lean_ctor_get(x_13, 11); +x_38 = lean_ctor_get(x_13, 12); +x_39 = lean_ctor_get(x_13, 13); +x_40 = lean_ctor_get(x_13, 14); +x_41 = lean_ctor_get(x_13, 15); +x_42 = lean_ctor_get(x_13, 16); +x_43 = lean_ctor_get(x_13, 17); +x_44 = lean_ctor_get(x_13, 18); lean_inc(x_44); lean_inc(x_43); lean_inc(x_42); lean_inc(x_41); lean_inc(x_40); lean_inc(x_39); -lean_dec(x_27); -x_54 = lean_box(0); -x_55 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_26, x_1, x_54); -x_56 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_41, x_21, x_55); -x_57 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_57, 0, x_39); -lean_ctor_set(x_57, 1, x_40); -lean_ctor_set(x_57, 2, x_56); -lean_ctor_set(x_57, 3, x_42); -lean_ctor_set(x_57, 4, x_43); -lean_ctor_set(x_57, 5, x_44); -lean_ctor_set(x_57, 6, x_46); -lean_ctor_set(x_57, 7, x_47); -lean_ctor_set(x_57, 8, x_48); -lean_ctor_set(x_57, 9, x_49); -lean_ctor_set(x_57, 10, x_50); -lean_ctor_set(x_57, 11, x_51); -lean_ctor_set(x_57, 12, x_52); -lean_ctor_set(x_57, 13, x_53); -lean_ctor_set_uint8(x_57, sizeof(void*)*14, x_45); -x_58 = lean_st_ref_set(x_3, x_57, x_28); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - lean_ctor_release(x_58, 1); - x_60 = x_58; +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_13); +x_45 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_26, x_1, x_2); +x_46 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_46, 0, x_25); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_46, 2, x_27); +lean_ctor_set(x_46, 3, x_28); +lean_ctor_set(x_46, 4, x_29); +lean_ctor_set(x_46, 5, x_30); +lean_ctor_set(x_46, 6, x_32); +lean_ctor_set(x_46, 7, x_33); +lean_ctor_set(x_46, 8, x_34); +lean_ctor_set(x_46, 9, x_35); +lean_ctor_set(x_46, 10, x_36); +lean_ctor_set(x_46, 11, x_37); +lean_ctor_set(x_46, 12, x_38); +lean_ctor_set(x_46, 13, x_39); +lean_ctor_set(x_46, 14, x_40); +lean_ctor_set(x_46, 15, x_41); +lean_ctor_set(x_46, 16, x_42); +lean_ctor_set(x_46, 17, x_43); +lean_ctor_set(x_46, 18, x_44); +lean_ctor_set_uint8(x_46, sizeof(void*)*19, x_31); +x_47 = lean_st_ref_set(x_3, x_46, x_14); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; } else { - lean_dec_ref(x_58); - x_60 = lean_box(0); + lean_dec_ref(x_47); + x_49 = lean_box(0); } -if (lean_is_scalar(x_60)) { - x_61 = lean_alloc_ctor(0, 2, 0); +x_50 = lean_box(0); +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(0, 2, 0); } else { - x_61 = x_60; -} -lean_ctor_set(x_61, 0, x_54); -lean_ctor_set(x_61, 1, x_59); -return x_61; -} + x_51 = x_49; } +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { size_t x_7; lean_object* x_8; x_7 = lean_unbox_usize(x_1); lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_registerParent___spec__5(x_7, x_2, x_3, x_4, x_5, x_6); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_7, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { size_t x_6; size_t x_7; lean_object* x_8; @@ -14399,46 +19225,15 @@ x_6 = lean_unbox_usize(x_2); lean_dec(x_2); x_7 = lean_unbox_usize(x_3); lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_registerParent___spec__4(x_1, x_6, x_7, x_4, x_5); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_6, x_7, x_4, x_5); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Grind_registerParent___spec__9(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Grind_registerParent___spec__8(x_1, x_4, x_3); -lean_dec(x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_registerParent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_Meta_Grind_registerParent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_Meta_Grind_setENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -14447,597 +19242,716 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_2, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_14 = lean_st_ref_get(x_5, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_st_ref_get(x_5, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +x_21 = l_Lean_Expr_isLambda(x_1); +x_22 = lean_ctor_get(x_18, 7); +lean_inc(x_22); +lean_dec(x_18); +x_23 = lean_ctor_get(x_15, 6); +lean_inc(x_23); +lean_dec(x_15); +x_24 = 0; +x_25 = lean_unsigned_to_nat(1u); +lean_inc_n(x_1, 4); +x_26 = lean_alloc_ctor(0, 10, 5); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_1); +lean_ctor_set(x_26, 2, x_1); +lean_ctor_set(x_26, 3, x_1); +lean_ctor_set(x_26, 4, x_20); +lean_ctor_set(x_26, 5, x_20); +lean_ctor_set(x_26, 6, x_25); +lean_ctor_set(x_26, 7, x_22); +lean_ctor_set(x_26, 8, x_4); +lean_ctor_set(x_26, 9, x_23); +lean_ctor_set_uint8(x_26, sizeof(void*)*10, x_24); +lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 1, x_2); +lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 2, x_3); +lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 3, x_21); +lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 4, x_24); +x_27 = l_Lean_Meta_Grind_setENode(x_1, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_st_ref_take(x_5, x_28); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = !lean_is_exclusive(x_30); +if (x_32 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 2); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_14, x_1); -if (lean_obj_tag(x_15) == 0) +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_30, 7); +x_34 = lean_nat_add(x_33, x_25); +lean_dec(x_33); +lean_ctor_set(x_30, 7, x_34); +x_35 = lean_st_ref_set(x_5, x_30, x_31); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_16; -x_16 = lean_box(0); -lean_ctor_set(x_11, 0, x_16); -return x_11; +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +lean_ctor_set(x_35, 0, x_38); +return x_35; } else { -lean_object* x_17; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -lean_ctor_set(x_11, 0, x_17); -return x_11; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_11); -x_20 = lean_ctor_get(x_18, 2); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_20, x_1); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_19); -return x_23; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_42 = lean_ctor_get(x_30, 0); +x_43 = lean_ctor_get(x_30, 1); +x_44 = lean_ctor_get(x_30, 2); +x_45 = lean_ctor_get(x_30, 3); +x_46 = lean_ctor_get(x_30, 4); +x_47 = lean_ctor_get(x_30, 5); +x_48 = lean_ctor_get_uint8(x_30, sizeof(void*)*19); +x_49 = lean_ctor_get(x_30, 6); +x_50 = lean_ctor_get(x_30, 7); +x_51 = lean_ctor_get(x_30, 8); +x_52 = lean_ctor_get(x_30, 9); +x_53 = lean_ctor_get(x_30, 10); +x_54 = lean_ctor_get(x_30, 11); +x_55 = lean_ctor_get(x_30, 12); +x_56 = lean_ctor_get(x_30, 13); +x_57 = lean_ctor_get(x_30, 14); +x_58 = lean_ctor_get(x_30, 15); +x_59 = lean_ctor_get(x_30, 16); +x_60 = lean_ctor_get(x_30, 17); +x_61 = lean_ctor_get(x_30, 18); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_30); +x_62 = lean_nat_add(x_50, x_25); +lean_dec(x_50); +x_63 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_63, 0, x_42); +lean_ctor_set(x_63, 1, x_43); +lean_ctor_set(x_63, 2, x_44); +lean_ctor_set(x_63, 3, x_45); +lean_ctor_set(x_63, 4, x_46); +lean_ctor_set(x_63, 5, x_47); +lean_ctor_set(x_63, 6, x_49); +lean_ctor_set(x_63, 7, x_62); +lean_ctor_set(x_63, 8, x_51); +lean_ctor_set(x_63, 9, x_52); +lean_ctor_set(x_63, 10, x_53); +lean_ctor_set(x_63, 11, x_54); +lean_ctor_set(x_63, 12, x_55); +lean_ctor_set(x_63, 13, x_56); +lean_ctor_set(x_63, 14, x_57); +lean_ctor_set(x_63, 15, x_58); +lean_ctor_set(x_63, 16, x_59); +lean_ctor_set(x_63, 17, x_60); +lean_ctor_set(x_63, 18, x_61); +lean_ctor_set_uint8(x_63, sizeof(void*)*19, x_48); +x_64 = lean_st_ref_set(x_5, x_63, x_31); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_66 = x_64; +} else { + lean_dec_ref(x_64); + x_66 = lean_box(0); } -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 0); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_19); -return x_25; +x_67 = lean_box(0); +if (lean_is_scalar(x_66)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_66; } +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_65); +return x_68; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParents___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getParents(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_unbox(x_2); +lean_dec(x_2); +x_15 = lean_unbox(x_3); +lean_dec(x_3); +x_16 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_array_get_size(x_1); -x_5 = lean_nat_dec_lt(x_3, x_4); -lean_dec(x_4); -if (x_5 == 0) -{ -lean_object* x_6; -lean_dec(x_3); -x_6 = lean_box(0); -return x_6; -} -else -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_fget(x_1, x_3); -x_8 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_7, x_2); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_unsigned_to_nat(1u); -x_10 = lean_nat_add(x_3, x_9); -lean_dec(x_3); -x_3 = x_10; -goto _start; -} -else -{ -lean_object* x_12; -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_3); -return x_12; -} -} +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Lean_Meta_isConstructorAppCore_x3f(x_1, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_16 = l_Lean_Meta_Grind_isInterpreted(x_1, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_7 = lean_usize_land(x_2, x_6); -x_8 = lean_usize_to_nat(x_7); -x_9 = lean_box(2); -x_10 = lean_array_get(x_9, x_4, x_8); -switch (lean_obj_tag(x_10)) { -case 0: +if (lean_obj_tag(x_16) == 0) { -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); +lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 0; +x_20 = lean_unbox(x_17); +lean_dec(x_17); +x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_19, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); lean_dec(x_11); -if (x_12 == 0) -{ +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_4); -return x_1; +return x_21; } else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_1); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_1, 0); -lean_dec(x_14); -x_15 = lean_array_set(x_4, x_8, x_9); +uint8_t x_22; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_15); -return x_1; +lean_dec(x_2); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; } else { -lean_object* x_16; lean_object* x_17; -lean_dec(x_1); -x_16 = lean_array_set(x_4, x_8, x_9); -lean_dec(x_8); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } -case 1: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_1); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_1, 0); -lean_dec(x_19); -x_20 = !lean_is_exclusive(x_10); -if (x_20 == 0) +else { -lean_object* x_21; lean_object* x_22; size_t x_23; lean_object* x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_10, 0); -x_22 = lean_array_set(x_4, x_8, x_9); -x_23 = lean_usize_shift_right(x_2, x_5); -x_24 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_21, x_23, x_3); -lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_24); -if (lean_obj_tag(x_25) == 0) +lean_dec(x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_26; -lean_ctor_set(x_10, 0, x_24); -x_26 = lean_array_set(x_22, x_8, x_10); +lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_16, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_16, 1); +lean_inc(x_27); +lean_dec(x_16); +x_28 = 1; +x_29 = lean_unbox(x_26); +lean_dec(x_26); +x_30 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_29, x_28, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_26); -return x_1; +return x_30; } else { -lean_object* x_27; uint8_t x_28; -lean_dec(x_24); -lean_free_object(x_10); -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; -x_29 = lean_array_set(x_22, x_8, x_27); +uint8_t x_31; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_29); -return x_1; +lean_dec(x_2); +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_16); +if (x_31 == 0) +{ +return x_16; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_27, 0); -x_31 = lean_ctor_get(x_27, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_27); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_array_set(x_22, x_8, x_32); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_33); -return x_1; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_16, 0); +x_33 = lean_ctor_get(x_16, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_16); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Meta_Grind_alreadyInternalized(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l_Lean_Meta_Grind_mkENode___lambda__1(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_17; +} else { -lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; -x_34 = lean_ctor_get(x_10, 0); -lean_inc(x_34); +uint8_t x_18; lean_dec(x_10); -x_35 = lean_array_set(x_4, x_8, x_9); -x_36 = lean_usize_shift_right(x_2, x_5); -x_37 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_34, x_36, x_3); -lean_inc(x_37); -x_38 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_37); -x_40 = lean_array_set(x_35, x_8, x_39); +lean_dec(x_9); lean_dec(x_8); -lean_ctor_set(x_1, 0, x_40); -return x_1; +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_12, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_12, 0, x_20); +return x_12; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -lean_dec(x_37); -x_41 = lean_ctor_get(x_38, 0); -lean_inc(x_41); -lean_dec(x_38); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - x_44 = x_41; -} else { - lean_dec_ref(x_41); - x_44 = lean_box(0); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; } -if (lean_is_scalar(x_44)) { - x_45 = lean_alloc_ctor(0, 2, 0); -} else { - x_45 = x_44; } -lean_ctor_set(x_45, 0, x_42); -lean_ctor_set(x_45, 1, x_43); -x_46 = lean_array_set(x_35, x_8, x_45); -lean_dec(x_8); -lean_ctor_set(x_1, 0, x_46); -return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_mkENode___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_13; +} } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; lean_object* x_51; lean_object* x_52; -lean_dec(x_1); -x_47 = lean_ctor_get(x_10, 0); -lean_inc(x_47); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - x_48 = x_10; -} else { - lean_dec_ref(x_10); - x_48 = lean_box(0); +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; } -x_49 = lean_array_set(x_4, x_8, x_9); -x_50 = lean_usize_shift_right(x_2, x_5); -x_51 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_47, x_50, x_3); -lean_inc(x_51); -x_52 = l_Lean_PersistentHashMap_isUnaryNode___rarg(x_51); -if (lean_obj_tag(x_52) == 0) +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -if (lean_is_scalar(x_48)) { - x_53 = lean_alloc_ctor(1, 1, 0); -} else { - x_53 = x_48; +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_13); +lean_dec(x_13); +x_15 = lean_box(x_14); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_11, 0); +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_11); +x_18 = l_Lean_Meta_Grind_ENode_isCongrRoot(x_16); +lean_dec(x_16); +x_19 = lean_box(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +return x_20; } -lean_ctor_set(x_53, 0, x_51); -x_54 = lean_array_set(x_49, x_8, x_53); -lean_dec(x_8); -x_55 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_55, 0, x_54); -return x_55; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_51); -lean_dec(x_48); -x_56 = lean_ctor_get(x_52, 0); -lean_inc(x_56); -lean_dec(x_52); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_59 = x_56; -} else { - lean_dec_ref(x_56); - x_59 = lean_box(0); +uint8_t x_21; +x_21 = !lean_is_exclusive(x_11); +if (x_21 == 0) +{ +return x_11; } -if (lean_is_scalar(x_59)) { - x_60 = lean_alloc_ctor(0, 2, 0); -} else { - x_60 = x_59; +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_11, 0); +x_23 = lean_ctor_get(x_11, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_11); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; } -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_58); -x_61 = lean_array_set(x_49, x_8, x_60); -lean_dec(x_8); -x_62 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_62, 0, x_61); -return x_62; } } } -default: +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isCongrRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -return x_1; +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 3); +lean_inc(x_12); +lean_dec(x_1); +x_13 = l_Lean_Meta_Grind_getCongrRoot(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = lean_ctor_get(x_13, 3); +lean_inc(x_15); +x_16 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_15, x_1); +lean_dec(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_free_object(x_11); +lean_dec(x_1); +x_17 = lean_box(0); +x_18 = l_Lean_Meta_Grind_getCongrRoot___lambda__1(x_13, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +return x_18; } +else +{ +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_1); +return x_11; } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = lean_ctor_get(x_1, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_1, 1); -lean_inc(x_64); -x_65 = lean_unsigned_to_nat(0u); -x_66 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_63, x_3, x_65); -if (lean_obj_tag(x_66) == 0) +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_11, 0); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_11); +x_21 = lean_ctor_get(x_19, 3); +lean_inc(x_21); +x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_21, x_1); +lean_dec(x_21); +if (x_22 == 0) { -lean_dec(x_64); -lean_dec(x_63); -return x_1; +lean_object* x_23; lean_object* x_24; +lean_dec(x_1); +x_23 = lean_box(0); +x_24 = l_Lean_Meta_Grind_getCongrRoot___lambda__1(x_19, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +return x_24; } else { -uint8_t x_67; -x_67 = !lean_is_exclusive(x_1); -if (x_67 == 0) -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_68 = lean_ctor_get(x_1, 1); -lean_dec(x_68); -x_69 = lean_ctor_get(x_1, 0); -lean_dec(x_69); -x_70 = lean_ctor_get(x_66, 0); -lean_inc(x_70); -lean_dec(x_66); -lean_inc(x_70); -x_71 = l_Array_eraseIdx___rarg(x_63, x_70, lean_box(0)); -x_72 = l_Array_eraseIdx___rarg(x_64, x_70, lean_box(0)); -lean_ctor_set(x_1, 1, x_72); -lean_ctor_set(x_1, 0, x_71); -return x_1; +lean_object* x_25; +lean_dec(x_19); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_20); +return x_25; +} +} } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +uint8_t x_26; lean_dec(x_1); -x_73 = lean_ctor_get(x_66, 0); -lean_inc(x_73); -lean_dec(x_66); -lean_inc(x_73); -x_74 = l_Array_eraseIdx___rarg(x_63, x_73, lean_box(0)); -x_75 = l_Array_eraseIdx___rarg(x_64, x_73, lean_box(0)); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; +x_26 = !lean_is_exclusive(x_11); +if (x_26 == 0) +{ +return x_11; } +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_11, 0); +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_11); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint64_t x_3; size_t x_4; lean_object* x_5; -x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_2); -return x_5; +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_getCongrRoot___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getCongrRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_11 = l_Lean_Meta_Grind_getParents(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_st_ref_take(x_2, x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getCongrRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_15, 2); -x_19 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_18, x_1); -lean_ctor_set(x_15, 2, x_19); -x_20 = lean_st_ref_set(x_2, x_15, x_16); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_22; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -lean_ctor_set(x_20, 0, x_12); -return x_20; +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*19); +lean_dec(x_12); +x_14 = lean_box(x_13); +lean_ctor_set(x_10, 0, x_14); +return x_10; } else { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_12); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_10); +x_17 = lean_ctor_get_uint8(x_15, sizeof(void*)*19); +lean_dec(x_15); +x_18 = lean_box(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +return x_19; } } -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_25 = lean_ctor_get(x_15, 0); -x_26 = lean_ctor_get(x_15, 1); -x_27 = lean_ctor_get(x_15, 2); -x_28 = lean_ctor_get(x_15, 3); -x_29 = lean_ctor_get(x_15, 4); -x_30 = lean_ctor_get(x_15, 5); -x_31 = lean_ctor_get_uint8(x_15, sizeof(void*)*14); -x_32 = lean_ctor_get(x_15, 6); -x_33 = lean_ctor_get(x_15, 7); -x_34 = lean_ctor_get(x_15, 8); -x_35 = lean_ctor_get(x_15, 9); -x_36 = lean_ctor_get(x_15, 10); -x_37 = lean_ctor_get(x_15, 11); -x_38 = lean_ctor_get(x_15, 12); -x_39 = lean_ctor_get(x_15, 13); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_15); -x_40 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_27, x_1); -x_41 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_26); -lean_ctor_set(x_41, 2, x_40); -lean_ctor_set(x_41, 3, x_28); -lean_ctor_set(x_41, 4, x_29); -lean_ctor_set(x_41, 5, x_30); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_33); -lean_ctor_set(x_41, 8, x_34); -lean_ctor_set(x_41, 9, x_35); -lean_ctor_set(x_41, 10, x_36); -lean_ctor_set(x_41, 11, x_37); -lean_ctor_set(x_41, 12, x_38); -lean_ctor_set(x_41, 13, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); -x_42 = lean_st_ref_set(x_2, x_41, x_16); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); } -if (lean_is_scalar(x_44)) { - x_45 = lean_alloc_ctor(0, 2, 0); -} else { - x_45 = x_44; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Grind_isInconsistent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; } -lean_ctor_set(x_45, 0, x_12); -lean_ctor_set(x_45, 1, x_43); -return x_45; } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_4; -x_4 = l_Array_indexOfAux___at_Lean_Meta_Grind_getParentsAndReset___spec__3(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_4; +lean_object* x_12; +x_12 = lean_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_eraseAux___at_Lean_Meta_Grind_getParentsAndReset___spec__2(x_1, x_4, x_3); -lean_dec(x_3); -return x_5; -} +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +lean_inc(x_1); +x_12 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +return x_16; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_erase___at_Lean_Meta_Grind_getParentsAndReset___spec__1(x_1, x_2); -lean_dec(x_2); -return x_3; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getParentsAndReset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getParentsAndReset(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_19; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -15047,927 +19961,1035 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -if (lean_obj_tag(x_1) == 0) +x_19 = !lean_is_exclusive(x_12); +if (x_19 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_2); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -return x_13; +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_1, 1); -lean_inc(x_15); -x_16 = lean_ctor_get(x_1, 3); -lean_inc(x_16); -lean_dec(x_1); -x_17 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_18, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_12, 0); +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_box(0); -x_22 = l_Lean_RBNode_insert___at_Lean_Meta_Grind_registerParent___spec__1(x_20, x_15, x_21); -x_1 = x_16; -x_2 = x_22; -x_11 = x_19; -goto _start; +lean_dec(x_12); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_54; -x_12 = lean_st_ref_get(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 2); -lean_inc(x_15); -lean_dec(x_13); -x_54 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Grind_registerParent___spec__7(x_15, x_2); -if (lean_obj_tag(x_54) == 0) -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_55 = lean_box(0); -x_56 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_55, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = lean_ctor_get(x_57, 0); -lean_inc(x_59); -lean_dec(x_57); -x_16 = x_59; -x_17 = x_58; -goto block_53; +lean_dec(x_11); +x_14 = lean_grind_mk_eq_proof(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_14; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_60 = lean_ctor_get(x_54, 0); -lean_inc(x_60); -lean_dec(x_54); -x_61 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_60, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_ctor_get(x_62, 0); -lean_inc(x_64); -lean_dec(x_62); -x_16 = x_64; -x_17 = x_63; -goto block_53; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_grind_mk_eq_proof(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_14; } -block_53: +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_st_ref_take(x_3, x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_st_ref_take(x_2, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_ctor_get(x_19, 2); -x_23 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_22, x_2, x_16); -lean_ctor_set(x_19, 2, x_23); -x_24 = lean_st_ref_set(x_3, x_19, x_20); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +uint8_t x_15; lean_object* x_16; uint8_t x_17; +x_15 = 1; +lean_ctor_set_uint8(x_12, sizeof(void*)*19, x_15); +x_16 = lean_st_ref_set(x_2, x_12, x_13); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_24, 0, x_27); -return x_24; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_16, 0, x_19); +return x_16; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_24, 1); -lean_inc(x_28); -lean_dec(x_24); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_31 = lean_ctor_get(x_19, 0); -x_32 = lean_ctor_get(x_19, 1); -x_33 = lean_ctor_get(x_19, 2); -x_34 = lean_ctor_get(x_19, 3); -x_35 = lean_ctor_get(x_19, 4); -x_36 = lean_ctor_get(x_19, 5); -x_37 = lean_ctor_get_uint8(x_19, sizeof(void*)*14); -x_38 = lean_ctor_get(x_19, 6); -x_39 = lean_ctor_get(x_19, 7); -x_40 = lean_ctor_get(x_19, 8); -x_41 = lean_ctor_get(x_19, 9); -x_42 = lean_ctor_get(x_19, 10); -x_43 = lean_ctor_get(x_19, 11); -x_44 = lean_ctor_get(x_19, 12); -x_45 = lean_ctor_get(x_19, 13); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; 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_23 = lean_ctor_get(x_12, 0); +x_24 = lean_ctor_get(x_12, 1); +x_25 = lean_ctor_get(x_12, 2); +x_26 = lean_ctor_get(x_12, 3); +x_27 = lean_ctor_get(x_12, 4); +x_28 = lean_ctor_get(x_12, 5); +x_29 = lean_ctor_get(x_12, 6); +x_30 = lean_ctor_get(x_12, 7); +x_31 = lean_ctor_get(x_12, 8); +x_32 = lean_ctor_get(x_12, 9); +x_33 = lean_ctor_get(x_12, 10); +x_34 = lean_ctor_get(x_12, 11); +x_35 = lean_ctor_get(x_12, 12); +x_36 = lean_ctor_get(x_12, 13); +x_37 = lean_ctor_get(x_12, 14); +x_38 = lean_ctor_get(x_12, 15); +x_39 = lean_ctor_get(x_12, 16); +x_40 = lean_ctor_get(x_12, 17); +x_41 = lean_ctor_get(x_12, 18); lean_inc(x_41); lean_inc(x_40); lean_inc(x_39); lean_inc(x_38); +lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); -lean_dec(x_19); -x_46 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_registerParent___spec__3(x_33, x_2, x_16); -x_47 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_47, 0, x_31); -lean_ctor_set(x_47, 1, x_32); -lean_ctor_set(x_47, 2, x_46); -lean_ctor_set(x_47, 3, x_34); -lean_ctor_set(x_47, 4, x_35); -lean_ctor_set(x_47, 5, x_36); -lean_ctor_set(x_47, 6, x_38); -lean_ctor_set(x_47, 7, x_39); -lean_ctor_set(x_47, 8, x_40); -lean_ctor_set(x_47, 9, x_41); -lean_ctor_set(x_47, 10, x_42); -lean_ctor_set(x_47, 11, x_43); -lean_ctor_set(x_47, 12, x_44); -lean_ctor_set(x_47, 13, x_45); -lean_ctor_set_uint8(x_47, sizeof(void*)*14, x_37); -x_48 = lean_st_ref_set(x_3, x_47, x_20); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -if (lean_is_exclusive(x_48)) { - lean_ctor_release(x_48, 0); - lean_ctor_release(x_48, 1); - x_50 = x_48; -} else { - lean_dec_ref(x_48); - x_50 = lean_box(0); -} -x_51 = lean_box(0); -if (lean_is_scalar(x_50)) { - x_52 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_12); +x_42 = 1; +x_43 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_43, 0, x_23); +lean_ctor_set(x_43, 1, x_24); +lean_ctor_set(x_43, 2, x_25); +lean_ctor_set(x_43, 3, x_26); +lean_ctor_set(x_43, 4, x_27); +lean_ctor_set(x_43, 5, x_28); +lean_ctor_set(x_43, 6, x_29); +lean_ctor_set(x_43, 7, x_30); +lean_ctor_set(x_43, 8, x_31); +lean_ctor_set(x_43, 9, x_32); +lean_ctor_set(x_43, 10, x_33); +lean_ctor_set(x_43, 11, x_34); +lean_ctor_set(x_43, 12, x_35); +lean_ctor_set(x_43, 13, x_36); +lean_ctor_set(x_43, 14, x_37); +lean_ctor_set(x_43, 15, x_38); +lean_ctor_set(x_43, 16, x_39); +lean_ctor_set(x_43, 17, x_40); +lean_ctor_set(x_43, 18, x_41); +lean_ctor_set_uint8(x_43, sizeof(void*)*19, x_42); +x_44 = lean_st_ref_set(x_2, x_43, x_13); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_46 = x_44; } else { - x_52 = x_50; -} -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_49); -return x_52; -} + lean_dec_ref(x_44); + x_46 = lean_box(0); } +x_47 = lean_box(0); +if (lean_is_scalar(x_46)) { + x_48 = lean_alloc_ctor(0, 2, 0); +} else { + x_48 = x_46; } +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_45); +return x_48; } -LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_RBNode_forIn_visit___at_Lean_Meta_Grind_copyParentsTo___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_copyParentsTo___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_copyParentsTo(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_Grind_markAsInconsistent___closed__1() { _start: { -lean_object* x_2; -x_2 = lean_ctor_get(x_1, 3); -lean_inc(x_2); -return x_2; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_markAsInconsistent___lambda__1___boxed), 10, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_Grind_markAsInconsistent___closed__2() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed), 1, 0); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("closed `", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___rarg___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_Grind_markAsInconsistent___closed__3() { _start: { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_setENode_unsafe__1___rarg(x_1); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_markAsInconsistent___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode_unsafe__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_3; -x_3 = l_Lean_Meta_Grind_setENode_unsafe__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get_uint8(x_11, sizeof(void*)*19); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l_Lean_Meta_Grind_updateLastTag___closed__1; +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_14, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = l_Lean_Meta_Grind_markAsInconsistent___closed__1; +x_20 = lean_unbox(x_17); +lean_dec(x_17); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_free_object(x_15); +x_21 = lean_box(0); +x_22 = lean_apply_10(x_19, x_21, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_22; } +else +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_st_ref_get(x_1, x_18); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_MVarId_getTag(x_27, x_5, x_6, x_7, x_8, x_26); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_MessageData_ofName(x_29); +x_32 = l_Lean_Meta_Grind_markAsInconsistent___closed__3; +lean_ctor_set_tag(x_23, 7); +lean_ctor_set(x_23, 1, x_31); +lean_ctor_set(x_23, 0, x_32); +x_33 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_15, 7); +lean_ctor_set(x_15, 1, x_33); +lean_ctor_set(x_15, 0, x_23); +x_34 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_14, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_30); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_apply_10(x_19, x_35, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_36); +return x_37; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); +uint8_t x_38; +lean_free_object(x_23); +lean_free_object(x_15); +lean_dec(x_8); lean_dec(x_7); -if (x_8 == 0) -{ +lean_dec(x_6); lean_dec(x_5); -return x_6; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_28); +if (x_38 == 0) +{ +return x_28; } else { -lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_9); -x_12 = lean_uint64_to_usize(x_11); -x_13 = 1; -x_14 = lean_usize_sub(x_1, x_13); -x_15 = 5; -x_16 = lean_usize_mul(x_15, x_14); -x_17 = lean_usize_shift_right(x_12, x_16); -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_5, x_18); -lean_dec(x_5); -x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_6, x_17, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_19; -x_6 = x_20; -goto _start; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_28, 0); +x_40 = lean_ctor_get(x_28, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_28); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) +else { -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_23, 0); +x_43 = lean_ctor_get(x_23, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_23); +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Lean_MVarId_getTag(x_44, x_5, x_6, x_7, x_8, x_43); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_1, 1); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 0); -lean_dec(x_11); -x_12 = lean_array_push(x_5, x_3); -x_13 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_12); -return x_1; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_MessageData_ofName(x_46); +x_49 = l_Lean_Meta_Grind_markAsInconsistent___closed__3; +x_50 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_48); +x_51 = l_Lean_Meta_Grind_updateLastTag___closed__5; +lean_ctor_set_tag(x_15, 7); +lean_ctor_set(x_15, 1, x_51); +lean_ctor_set(x_15, 0, x_50); +x_52 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_14, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_47); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_apply_10(x_19, x_53, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_54); +return x_55; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_free_object(x_15); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_14 = lean_array_push(x_5, x_3); -x_15 = lean_array_push(x_6, x_4); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; +x_56 = lean_ctor_get(x_45, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_45, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_58 = x_45; +} else { + lean_dec_ref(x_45); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} +} } } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_5, x_2); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_15, 0); +x_61 = lean_ctor_get(x_15, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_15); +x_62 = l_Lean_Meta_Grind_markAsInconsistent___closed__1; +x_63 = lean_unbox(x_60); +lean_dec(x_60); +if (x_63 == 0) { -lean_object* x_19; lean_object* x_20; -lean_dec(x_6); -lean_dec(x_5); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_2, x_19); -lean_dec(x_2); -x_2 = x_20; -goto _start; +lean_object* x_64; lean_object* x_65; +x_64 = lean_box(0); +x_65 = lean_apply_10(x_62, x_64, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_61); +return x_65; } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -if (x_22 == 0) +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_66 = lean_st_ref_get(x_1, x_61); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_69 = x_66; +} else { + lean_dec_ref(x_66); + x_69 = lean_box(0); +} +x_70 = lean_ctor_get(x_67, 0); +lean_inc(x_70); +lean_dec(x_67); +x_71 = l_Lean_MVarId_getTag(x_70, x_5, x_6, x_7, x_8, x_68); +if (lean_obj_tag(x_71) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 1); -lean_dec(x_23); -x_24 = lean_ctor_get(x_1, 0); -lean_dec(x_24); -x_25 = lean_array_fset(x_5, x_2, x_3); -x_26 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -lean_ctor_set(x_1, 1, x_26); -lean_ctor_set(x_1, 0, x_25); -return x_1; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +x_74 = l_Lean_MessageData_ofName(x_72); +x_75 = l_Lean_Meta_Grind_markAsInconsistent___closed__3; +if (lean_is_scalar(x_69)) { + x_76 = lean_alloc_ctor(7, 2, 0); +} else { + x_76 = x_69; + lean_ctor_set_tag(x_76, 7); +} +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_74); +x_77 = l_Lean_Meta_Grind_updateLastTag___closed__5; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_14, x_78, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_73); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +lean_dec(x_79); +x_82 = lean_apply_10(x_62, x_80, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_81); +return x_82; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_1); -x_27 = lean_array_fset(x_5, x_2, x_3); -x_28 = lean_array_fset(x_6, x_2, x_4); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_69); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_dec(x_1); +x_83 = lean_ctor_get(x_71, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_71, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_85 = x_71; +} else { + lean_dec_ref(x_71); + x_85 = lean_box(0); } +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); +} else { + x_86 = x_85; } +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = 1; -x_9 = 5; -x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_14 = lean_nat_dec_lt(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) +} +else { -lean_dec(x_12); +uint8_t x_87; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_1; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_87 = !lean_is_exclusive(x_10); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_10, 0); +lean_dec(x_88); +x_89 = lean_box(0); +lean_ctor_set(x_10, 0, x_89); +return x_10; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_free_object(x_15); -x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_array_fset(x_17, x_12, x_23); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_24); -return x_1; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_10, 1); +lean_inc(x_90); +lean_dec(x_10); +x_91 = lean_box(0); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_90); +return x_92; } -else -{ -lean_object* x_25; -lean_dec(x_20); -lean_dec(x_19); -lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_4); -x_25 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_25); -return x_1; } } -else -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_26); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_array_fset(x_17, x_12, x_30); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_31); -return x_1; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_32; lean_object* x_33; -lean_dec(x_27); -lean_dec(x_26); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_4); -lean_ctor_set(x_32, 1, x_5); -x_33 = lean_array_fset(x_17, x_12, x_32); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_33); -return x_1; -} +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_markAsInconsistent___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; } } -case 1: +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_7, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -x_35 = lean_ctor_get(x_15, 0); -x_36 = lean_usize_shift_right(x_2, x_9); -x_37 = lean_usize_add(x_3, x_8); -x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_35, x_36, x_37, x_4, x_5); -lean_ctor_set(x_15, 0, x_38); -x_39 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); -return x_1; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_ctor_get(x_14, 7); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_15, x_1); +x_17 = lean_box(x_16); +lean_ctor_set(x_11, 0, x_17); +return x_11; } else { -lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_15, 0); -lean_inc(x_40); -lean_dec(x_15); -x_41 = lean_usize_shift_right(x_2, x_9); -x_42 = lean_usize_add(x_3, x_8); -x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_40, x_41, x_42, x_4, x_5); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_array_fset(x_17, x_12, x_44); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_45); -return x_1; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_11, 0); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_11); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_20, 7); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_21, x_1); +x_23 = lean_box(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_19); +return x_24; } } -default: +} +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_46; lean_object* x_47; -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_4); -lean_ctor_set(x_46, 1, x_5); -x_47 = lean_array_fset(x_17, x_12, x_46); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = lean_st_ref_take(x_8, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); lean_dec(x_12); -lean_ctor_set(x_1, 0, x_47); -return x_1; -} -} -} -} -else +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) { -lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_1, 0); -lean_inc(x_48); -lean_dec(x_1); -x_49 = 1; -x_50 = 5; -x_51 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; -x_52 = lean_usize_land(x_2, x_51); -x_53 = lean_usize_to_nat(x_52); -x_54 = lean_array_get_size(x_48); -x_55 = lean_nat_dec_lt(x_53, x_54); -lean_dec(x_54); -if (x_55 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) { -lean_object* x_56; -lean_dec(x_53); -lean_dec(x_5); -lean_dec(x_4); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_48); -return x_56; +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_14, 7); +x_20 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_19, x_1, x_2); +lean_ctor_set(x_14, 7, x_20); +x_21 = lean_st_ref_set(x_8, x_13, x_15); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_21, 0, x_24); +return x_21; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_array_fget(x_48, x_53); -x_58 = lean_box(0); -x_59 = lean_array_fset(x_48, x_53, x_58); -switch (lean_obj_tag(x_57)) { -case 0: -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_57, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_62 = x_57; -} else { - lean_dec_ref(x_57); - x_62 = lean_box(0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; } -x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_array_fset(x_59, x_53, x_65); -lean_dec(x_53); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_61); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_68 = lean_alloc_ctor(0, 2, 0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_28 = lean_ctor_get(x_14, 0); +x_29 = lean_ctor_get(x_14, 1); +x_30 = lean_ctor_get(x_14, 2); +x_31 = lean_ctor_get(x_14, 3); +x_32 = lean_ctor_get(x_14, 4); +x_33 = lean_ctor_get(x_14, 5); +x_34 = lean_ctor_get(x_14, 6); +x_35 = lean_ctor_get(x_14, 7); +x_36 = lean_ctor_get(x_14, 8); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_14); +x_37 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_35, x_1, x_2); +x_38 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_38, 0, x_28); +lean_ctor_set(x_38, 1, x_29); +lean_ctor_set(x_38, 2, x_30); +lean_ctor_set(x_38, 3, x_31); +lean_ctor_set(x_38, 4, x_32); +lean_ctor_set(x_38, 5, x_33); +lean_ctor_set(x_38, 6, x_34); +lean_ctor_set(x_38, 7, x_37); +lean_ctor_set(x_38, 8, x_36); +lean_ctor_set(x_13, 0, x_38); +x_39 = lean_st_ref_set(x_8, x_13, x_15); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; } else { - x_68 = x_62; -} -lean_ctor_set(x_68, 0, x_4); -lean_ctor_set(x_68, 1, x_5); -x_69 = lean_array_fset(x_59, x_53, x_68); -lean_dec(x_53); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; -} + lean_dec_ref(x_39); + x_41 = lean_box(0); } -case 1: -{ -lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_71 = lean_ctor_get(x_57, 0); -lean_inc(x_71); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - x_72 = x_57; +x_42 = lean_box(0); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_57); - x_72 = lean_box(0); + x_43 = x_41; } -x_73 = lean_usize_shift_right(x_2, x_50); -x_74 = lean_usize_add(x_3, x_49); -x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_71, x_73, x_74, x_4, x_5); -if (lean_is_scalar(x_72)) { - x_76 = lean_alloc_ctor(1, 1, 0); -} else { - x_76 = x_72; +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +return x_43; } -lean_ctor_set(x_76, 0, x_75); -x_77 = lean_array_fset(x_59, x_53, x_76); -lean_dec(x_53); -x_78 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_78, 0, x_77); -return x_78; } -default: +else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_4); -lean_ctor_set(x_79, 1, x_5); -x_80 = lean_array_fset(x_59, x_53, x_79); -lean_dec(x_53); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; -} -} -} +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_44 = lean_ctor_get(x_13, 1); +x_45 = lean_ctor_get(x_13, 2); +x_46 = lean_ctor_get(x_13, 3); +x_47 = lean_ctor_get(x_13, 4); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_13); +x_48 = lean_ctor_get(x_14, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_14, 1); +lean_inc(x_49); +x_50 = lean_ctor_get(x_14, 2); +lean_inc(x_50); +x_51 = lean_ctor_get(x_14, 3); +lean_inc(x_51); +x_52 = lean_ctor_get(x_14, 4); +lean_inc(x_52); +x_53 = lean_ctor_get(x_14, 5); +lean_inc(x_53); +x_54 = lean_ctor_get(x_14, 6); +lean_inc(x_54); +x_55 = lean_ctor_get(x_14, 7); +lean_inc(x_55); +x_56 = lean_ctor_get(x_14, 8); +lean_inc(x_56); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + lean_ctor_release(x_14, 2); + lean_ctor_release(x_14, 3); + lean_ctor_release(x_14, 4); + lean_ctor_release(x_14, 5); + lean_ctor_release(x_14, 6); + lean_ctor_release(x_14, 7); + lean_ctor_release(x_14, 8); + x_57 = x_14; +} else { + lean_dec_ref(x_14); + x_57 = lean_box(0); } +x_58 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_55, x_1, x_2); +if (lean_is_scalar(x_57)) { + x_59 = lean_alloc_ctor(0, 9, 0); +} else { + x_59 = x_57; } -else -{ -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) -{ -lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; -x_83 = lean_unsigned_to_nat(0u); -x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(x_1, x_83, x_4, x_5); -x_85 = 7; -x_86 = lean_usize_dec_le(x_85, x_3); -if (x_86 == 0) -{ -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_dec_lt(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_84, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 1); -lean_inc(x_91); -lean_dec(x_84); -x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_90, x_91, lean_box(0), x_83, x_92); -lean_dec(x_91); -lean_dec(x_90); -return x_93; +lean_ctor_set(x_59, 0, x_48); +lean_ctor_set(x_59, 1, x_49); +lean_ctor_set(x_59, 2, x_50); +lean_ctor_set(x_59, 3, x_51); +lean_ctor_set(x_59, 4, x_52); +lean_ctor_set(x_59, 5, x_53); +lean_ctor_set(x_59, 6, x_54); +lean_ctor_set(x_59, 7, x_58); +lean_ctor_set(x_59, 8, x_56); +x_60 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_44); +lean_ctor_set(x_60, 2, x_45); +lean_ctor_set(x_60, 3, x_46); +lean_ctor_set(x_60, 4, x_47); +x_61 = lean_st_ref_set(x_8, x_60, x_15); +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_63 = x_61; +} else { + lean_dec_ref(x_61); + x_63 = lean_box(0); } -else -{ -return x_84; +x_64 = lean_box(0); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_63; } +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; } -else -{ -return x_84; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; -x_94 = lean_ctor_get(x_1, 0); -x_95 = lean_ctor_get(x_1, 1); -lean_inc(x_95); -lean_inc(x_94); -lean_dec(x_1); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -x_97 = lean_unsigned_to_nat(0u); -x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_setENode___spec__4(x_96, x_97, x_4, x_5); -x_99 = 7; -x_100 = lean_usize_dec_le(x_99, x_3); -if (x_100 == 0) +lean_object* x_11; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_11 = l_Lean_Meta_Grind_markAsInconsistent(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); -x_102 = lean_unsigned_to_nat(4u); -x_103 = lean_nat_dec_lt(x_101, x_102); -lean_dec(x_101); -if (x_103 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_2, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_98, 1); -lean_inc(x_105); -lean_dec(x_98); -x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +lean_inc(x_16); +x_21 = l_Lean_MVarId_getType(x_16, x_6, x_7, x_8, x_9, x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_22); +x_24 = l_Lean_Expr_isFalse(x_22); +if (x_24 == 0) +{ +lean_object* x_25; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_25 = l_Lean_Meta_mkFalseElim(x_22, x_1, x_6, x_7, x_8, x_9, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_16, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_28; } else { -return x_98; -} +uint8_t x_29; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_29 = !lean_is_exclusive(x_25); +if (x_29 == 0) +{ +return x_25; } else { -return x_98; -} -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_25, 0); +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_25); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; -x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_5, x_6, x_2, x_3); -return x_7; -} +lean_object* x_33; +lean_dec(x_22); +x_33 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_16, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_33; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_st_ref_take(x_3, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_13, 1); -x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_16, x_1, x_2); -lean_ctor_set(x_13, 1, x_17); -x_18 = lean_st_ref_set(x_3, x_13, x_14); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_18, 0, x_21); -return x_18; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} +uint8_t x_34; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_34 = !lean_is_exclusive(x_21); +if (x_34 == 0) +{ +return x_21; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_25 = lean_ctor_get(x_13, 0); -x_26 = lean_ctor_get(x_13, 1); -x_27 = lean_ctor_get(x_13, 2); -x_28 = lean_ctor_get(x_13, 3); -x_29 = lean_ctor_get(x_13, 4); -x_30 = lean_ctor_get(x_13, 5); -x_31 = lean_ctor_get_uint8(x_13, sizeof(void*)*14); -x_32 = lean_ctor_get(x_13, 6); -x_33 = lean_ctor_get(x_13, 7); -x_34 = lean_ctor_get(x_13, 8); -x_35 = lean_ctor_get(x_13, 9); -x_36 = lean_ctor_get(x_13, 10); -x_37 = lean_ctor_get(x_13, 11); -x_38 = lean_ctor_get(x_13, 12); -x_39 = lean_ctor_get(x_13, 13); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_21, 0); +x_36 = lean_ctor_get(x_21, 1); lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_13); -x_40 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_setENode___spec__1(x_26, x_1, x_2); -x_41 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_41, 0, x_25); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_27); -lean_ctor_set(x_41, 3, x_28); -lean_ctor_set(x_41, 4, x_29); -lean_ctor_set(x_41, 5, x_30); -lean_ctor_set(x_41, 6, x_32); -lean_ctor_set(x_41, 7, x_33); -lean_ctor_set(x_41, 8, x_34); -lean_ctor_set(x_41, 9, x_35); -lean_ctor_set(x_41, 10, x_36); -lean_ctor_set(x_41, 11, x_37); -lean_ctor_set(x_41, 12, x_38); -lean_ctor_set(x_41, 13, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*14, x_31); -x_42 = lean_st_ref_set(x_3, x_41, x_14); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); +lean_dec(x_21); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} } -x_45 = lean_box(0); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; } -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +else +{ +uint8_t x_38; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_17); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_17, 0); +lean_dec(x_39); +x_40 = lean_box(0); +lean_ctor_set(x_17, 0, x_40); +return x_17; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_17, 1); +lean_inc(x_41); +lean_dec(x_17); +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +return x_43; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -size_t x_7; lean_object* x_8; -x_7 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_setENode___spec__3(x_7, x_2, x_3, x_4, x_5, x_6); +uint8_t x_44; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_8; +lean_dec(x_1); +x_44 = !lean_is_exclusive(x_11); +if (x_44 == 0) +{ +return x_11; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_11, 0); +x_46 = lean_ctor_get(x_11, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_11); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); +lean_object* x_11; +x_11 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_setENode___spec__2(x_1, x_6, x_7, x_4, x_5); -return x_8; +lean_dec(x_2); +lean_dec(x_1); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_setENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l_Lean_Meta_Grind_setENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -15979,505 +21001,305 @@ lean_dec(x_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore(lean_object* x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_14 = lean_st_ref_get(x_5, x_13); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_st_ref_get(x_5, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_box(0); -x_21 = l_Lean_Expr_isLambda(x_1); -x_22 = lean_ctor_get(x_18, 7); -lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_ctor_get(x_15, 6); -lean_inc(x_23); -lean_dec(x_15); -x_24 = 0; -x_25 = lean_unsigned_to_nat(1u); -lean_inc_n(x_1, 4); -x_26 = lean_alloc_ctor(0, 10, 5); -lean_ctor_set(x_26, 0, x_1); -lean_ctor_set(x_26, 1, x_1); -lean_ctor_set(x_26, 2, x_1); -lean_ctor_set(x_26, 3, x_1); -lean_ctor_set(x_26, 4, x_20); -lean_ctor_set(x_26, 5, x_20); -lean_ctor_set(x_26, 6, x_25); -lean_ctor_set(x_26, 7, x_22); -lean_ctor_set(x_26, 8, x_4); -lean_ctor_set(x_26, 9, x_23); -lean_ctor_set_uint8(x_26, sizeof(void*)*10, x_24); -lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 1, x_2); -lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 2, x_3); -lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 3, x_21); -lean_ctor_set_uint8(x_26, sizeof(void*)*10 + 4, x_24); -x_27 = l_Lean_Meta_Grind_setENode(x_1, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = lean_st_ref_take(x_5, x_28); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -x_32 = !lean_is_exclusive(x_30); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_33 = lean_ctor_get(x_30, 7); -x_34 = lean_nat_add(x_33, x_25); -lean_dec(x_33); -lean_ctor_set(x_30, 7, x_34); -x_35 = lean_st_ref_set(x_5, x_30, x_31); -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_35, 0); -lean_dec(x_37); -x_38 = lean_box(0); -lean_ctor_set(x_35, 0, x_38); -return x_35; +lean_object* x_4; lean_object* x_5; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +x_5 = lean_array_push(x_1, x_4); +return x_5; } -else +} +static lean_object* _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1() { +_start: { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_35, 1); -lean_inc(x_39); -lean_dec(x_35); -x_40 = lean_box(0); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1), 3, 0); +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(lean_object* x_1) { +_start: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_42 = lean_ctor_get(x_30, 0); -x_43 = lean_ctor_get(x_30, 1); -x_44 = lean_ctor_get(x_30, 2); -x_45 = lean_ctor_get(x_30, 3); -x_46 = lean_ctor_get(x_30, 4); -x_47 = lean_ctor_get(x_30, 5); -x_48 = lean_ctor_get_uint8(x_30, sizeof(void*)*14); -x_49 = lean_ctor_get(x_30, 6); -x_50 = lean_ctor_get(x_30, 7); -x_51 = lean_ctor_get(x_30, 8); -x_52 = lean_ctor_get(x_30, 9); -x_53 = lean_ctor_get(x_30, 10); -x_54 = lean_ctor_get(x_30, 11); -x_55 = lean_ctor_get(x_30, 12); -x_56 = lean_ctor_get(x_30, 13); -lean_inc(x_56); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_30); -x_57 = lean_nat_add(x_50, x_25); -lean_dec(x_50); -x_58 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_58, 0, x_42); -lean_ctor_set(x_58, 1, x_43); -lean_ctor_set(x_58, 2, x_44); -lean_ctor_set(x_58, 3, x_45); -lean_ctor_set(x_58, 4, x_46); -lean_ctor_set(x_58, 5, x_47); -lean_ctor_set(x_58, 6, x_49); -lean_ctor_set(x_58, 7, x_57); -lean_ctor_set(x_58, 8, x_51); -lean_ctor_set(x_58, 9, x_52); -lean_ctor_set(x_58, 10, x_53); -lean_ctor_set(x_58, 11, x_54); -lean_ctor_set(x_58, 12, x_55); -lean_ctor_set(x_58, 13, x_56); -lean_ctor_set_uint8(x_58, sizeof(void*)*14, x_48); -x_59 = lean_st_ref_set(x_5, x_58, x_31); -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_61 = x_59; -} else { - lean_dec_ref(x_59); - x_61 = lean_box(0); +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; +x_3 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_1, x_3); +return x_4; } -x_62 = lean_box(0); -if (lean_is_scalar(x_61)) { - x_63 = lean_alloc_ctor(0, 2, 0); -} else { - x_63 = x_61; } -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -return x_63; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENodeCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_14; uint8_t x_15; lean_object* x_16; -x_14 = lean_unbox(x_2); -lean_dec(x_2); -x_15 = lean_unbox(x_3); -lean_dec(x_3); -x_16 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_16; +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 7); +x_4 = lean_ctor_get(x_2, 7); +x_5 = lean_nat_dec_lt(x_3, x_4); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1() { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = l_Lean_Meta_isConstructorAppCore_x3f(x_1, x_8, x_9, x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_16 = l_Lean_Meta_Grind_isInterpreted(x_1, x_8, x_9, x_10, x_11, x_15); -if (lean_obj_tag(x_14) == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -if (lean_obj_tag(x_16) == 0) +uint8_t x_7; +x_7 = lean_nat_dec_lt(x_3, x_4); +if (x_7 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 0; -x_20 = lean_unbox(x_17); -lean_dec(x_17); -x_21 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_20, x_19, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_21; +lean_dec(x_3); +return x_2; } else { -uint8_t x_22; -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1; +lean_inc(x_3); +x_9 = l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(x_1, x_2, x_8, x_3, x_4, lean_box(0), lean_box(0)); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_16); -if (x_22 == 0) +x_12 = lean_nat_dec_le(x_4, x_10); +if (x_12 == 0) { -return x_16; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_1, x_11, x_3, x_10, lean_box(0), lean_box(0)); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_10, x_14); +lean_dec(x_10); +x_2 = x_13; +x_3 = x_15; +x_5 = lean_box(0); +x_6 = lean_box(0); +goto _start; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_16, 0); -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_16); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_dec(x_10); +lean_dec(x_3); +return x_11; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_dec(x_14); -if (lean_obj_tag(x_16) == 0) +lean_object* x_10; uint8_t x_11; +x_10 = lean_st_ref_get(x_1, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_26 = lean_ctor_get(x_16, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_16, 1); -lean_inc(x_27); -lean_dec(x_16); -x_28 = 1; -x_29 = lean_unbox(x_26); -lean_dec(x_26); -x_30 = l_Lean_Meta_Grind_mkENodeCore(x_1, x_29, x_28, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_27); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_30; -} -else +lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_13); +lean_dec(x_13); +x_15 = lean_array_size(x_14); +x_16 = 0; +x_17 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_15, x_16, x_14); +x_18 = lean_array_get_size(x_17); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_18, x_19); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_eq(x_18, x_21); +if (x_22 == 0) { -uint8_t x_31; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_2); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_16); -if (x_31 == 0) +uint8_t x_23; +x_23 = lean_nat_dec_le(x_21, x_20); +if (x_23 == 0) { -return x_16; +lean_object* x_24; +lean_inc(x_20); +x_24 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_18, x_17, x_20, x_20, lean_box(0), lean_box(0)); +lean_dec(x_20); +lean_dec(x_18); +lean_ctor_set(x_10, 0, x_24); +return x_10; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_16, 0); -x_33 = lean_ctor_get(x_16, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_16); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; -} -} -} +lean_object* x_25; +x_25 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_18, x_17, x_21, x_20, lean_box(0), lean_box(0)); +lean_dec(x_20); +lean_dec(x_18); +lean_ctor_set(x_10, 0, x_25); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = l_Lean_Meta_Grind_alreadyInternalized(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_box(0); -x_17 = l_Lean_Meta_Grind_mkENode___lambda__1(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_17; +lean_dec(x_20); +lean_dec(x_18); +lean_ctor_set(x_10, 0, x_17); +return x_10; +} } else { -uint8_t x_18; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_26 = lean_ctor_get(x_10, 0); +x_27 = lean_ctor_get(x_10, 1); +lean_inc(x_27); +lean_inc(x_26); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_12); -if (x_18 == 0) +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_28); +lean_dec(x_28); +x_30 = lean_array_size(x_29); +x_31 = 0; +x_32 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_30, x_31, x_29); +x_33 = lean_array_get_size(x_32); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_33, x_34); +x_36 = lean_unsigned_to_nat(0u); +x_37 = lean_nat_dec_eq(x_33, x_36); +if (x_37 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_12, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_12, 0, x_20); -return x_12; +uint8_t x_38; +x_38 = lean_nat_dec_le(x_36, x_35); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_inc(x_35); +x_39 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_33, x_32, x_35, x_35, lean_box(0), lean_box(0)); +lean_dec(x_35); +lean_dec(x_33); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_27); +return x_40; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} +lean_object* x_41; lean_object* x_42; +x_41 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_33, x_32, x_36, x_35, lean_box(0), lean_box(0)); +lean_dec(x_35); +lean_dec(x_33); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_27); +return x_42; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +else { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_mkENode___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_13; +lean_object* x_43; +lean_dec(x_35); +lean_dec(x_33); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_32); +lean_ctor_set(x_43, 1, x_27); +return x_43; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkENode___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Meta_Grind_mkENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; -lean_inc(x_1); -x_11 = l_Lean_Meta_Grind_getENode(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 3); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_14); -lean_dec(x_14); +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(x_1, x_2, x_3); lean_dec(x_1); -x_16 = lean_box(x_15); -lean_ctor_set(x_11, 0, x_16); -return x_11; +return x_4; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___boxed(lean_object* x_1) { +_start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_11); -x_19 = lean_ctor_get(x_17, 3); -lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_19); -lean_dec(x_19); +lean_object* x_2; +x_2 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_1); lean_dec(x_1); -x_21 = lean_box(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_18); -return x_22; +return x_2; } } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_23; +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); lean_dec(x_1); -x_23 = !lean_is_exclusive(x_11); -if (x_23 == 0) -{ -return x_11; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_11, 0); -x_25 = lean_ctor_get(x_11, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_11); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; -} -} +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_4, x_5, x_3); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isCongrRoot___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_isCongrRoot(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +uint8_t x_3; lean_object* x_4; +x_3 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(x_1, x_2); lean_dec(x_2); -return x_11; +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_ref_get(x_1, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*14); -lean_dec(x_12); -x_14 = lean_box(x_13); -lean_ctor_set(x_10, 0, x_14); -return x_10; -} -else -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_10, 0); -x_16 = lean_ctor_get(x_10, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_10); -x_17 = lean_ctor_get_uint8(x_15, sizeof(void*)*14); -lean_dec(x_15); -x_18 = lean_box(x_17); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_16); -return x_19; -} +lean_object* x_7; +x_7 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_4); +lean_dec(x_1); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isInconsistent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_Grind_isInconsistent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_getENodes(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -16489,154 +21311,203 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_12; -x_12 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkHEqProof___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_6, x_5); +if (x_17 == 0) { -lean_object* x_12; -x_12 = lean_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_12; -} +lean_object* x_18; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_16); +return x_18; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqHEqProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; +lean_object* x_19; lean_object* x_20; +lean_dec(x_7); +x_19 = lean_array_uget(x_4, x_6); +lean_inc(x_1); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_1); -x_12 = l_Lean_Meta_Grind_hasSameType(x_1, x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_grind_mk_heq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_16; -} -else +x_20 = lean_apply_10(x_1, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_grind_mk_eq_proof(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -return x_18; -} +lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = 1; +x_23 = lean_usize_add(x_6, x_22); +x_24 = lean_box(0); +x_6 = x_23; +x_7 = x_24; +x_16 = x_21; +goto _start; } else { -uint8_t x_19; +uint8_t x_26; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) +x_26 = !lean_is_exclusive(x_20); +if (x_26 == 0) { -return x_12; +return x_20; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_12, 0); -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_12); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_20, 0); +x_28 = lean_ctor_get(x_20, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_20); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqTrueProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = l_Lean_Meta_Grind_getTrueExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; +x_11 = l_Lean_Meta_Grind_getENodes(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = lean_grind_mk_eq_proof(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; +x_14 = lean_box(0); +x_15 = lean_array_size(x_12); +x_16 = 0; +x_17 = lean_box(0); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(x_1, x_12, x_14, x_12, x_15, x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_12); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_17); +return x_18; } +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_17); +lean_ctor_set(x_22, 1, x_21); +return x_22; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_mkEqFalseProof(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) +{ +return x_18; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_18, 0); +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_18); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = l_Lean_Meta_Grind_getFalseExpr___rarg(x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_grind_mk_eq_proof(x_1, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_18 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_st_ref_take(x_1, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) +lean_object* x_13; +lean_inc(x_3); +x_13 = lean_apply_10(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_14; lean_object* x_15; uint8_t x_16; -x_14 = 1; -lean_ctor_set_uint8(x_11, sizeof(void*)*14, x_14); -x_15 = lean_st_ref_set(x_1, x_11, x_12); -x_16 = !lean_is_exclusive(x_15); +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +uint8_t x_16; +lean_dec(x_3); +x_16 = !lean_is_exclusive(x_13); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_15, 0); +x_17 = lean_ctor_get(x_13, 0); lean_dec(x_17); x_18 = lean_box(0); -lean_ctor_set(x_15, 0, x_18); -return x_15; +lean_ctor_set(x_13, 0, x_18); +return x_13; } else { lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_15, 1); +x_19 = lean_ctor_get(x_13, 1); lean_inc(x_19); -lean_dec(x_15); +lean_dec(x_13); x_20 = lean_box(0); x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); @@ -16646,501 +21517,383 @@ return x_21; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_22 = lean_ctor_get(x_11, 0); -x_23 = lean_ctor_get(x_11, 1); -x_24 = lean_ctor_get(x_11, 2); -x_25 = lean_ctor_get(x_11, 3); -x_26 = lean_ctor_get(x_11, 4); -x_27 = lean_ctor_get(x_11, 5); -x_28 = lean_ctor_get(x_11, 6); -x_29 = lean_ctor_get(x_11, 7); -x_30 = lean_ctor_get(x_11, 8); -x_31 = lean_ctor_get(x_11, 9); -x_32 = lean_ctor_get(x_11, 10); -x_33 = lean_ctor_get(x_11, 11); -x_34 = lean_ctor_get(x_11, 12); -x_35 = lean_ctor_get(x_11, 13); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_22 = lean_ctor_get(x_13, 1); lean_inc(x_22); -lean_dec(x_11); -x_36 = 1; -x_37 = lean_alloc_ctor(0, 14, 1); -lean_ctor_set(x_37, 0, x_22); -lean_ctor_set(x_37, 1, x_23); -lean_ctor_set(x_37, 2, x_24); -lean_ctor_set(x_37, 3, x_25); -lean_ctor_set(x_37, 4, x_26); -lean_ctor_set(x_37, 5, x_27); -lean_ctor_set(x_37, 6, x_28); -lean_ctor_set(x_37, 7, x_29); -lean_ctor_set(x_37, 8, x_30); -lean_ctor_set(x_37, 9, x_31); -lean_ctor_set(x_37, 10, x_32); -lean_ctor_set(x_37, 11, x_33); -lean_ctor_set(x_37, 12, x_34); -lean_ctor_set(x_37, 13, x_35); -lean_ctor_set_uint8(x_37, sizeof(void*)*14, x_36); -x_38 = lean_st_ref_set(x_1, x_37, x_12); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_40 = x_38; -} else { - lean_dec_ref(x_38); - x_40 = lean_box(0); -} -x_41 = lean_box(0); -if (lean_is_scalar(x_40)) { - x_42 = lean_alloc_ctor(0, 2, 0); -} else { - x_42 = x_40; +lean_dec(x_13); +x_23 = lean_st_ref_take(x_2, x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_array_push(x_24, x_3); +x_27 = lean_st_ref_set(x_2, x_26, x_25); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +return x_27; } -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_39); -return x_42; +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_27, 0); +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_27); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markAsInconsistent___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_markAsInconsistent(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +uint8_t x_32; lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +x_32 = !lean_is_exclusive(x_13); +if (x_32 == 0) { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_7, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +return x_13; +} +else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_13, 0); +x_34 = lean_ctor_get(x_13, 1); +lean_inc(x_34); +lean_inc(x_33); lean_dec(x_13); -x_15 = lean_ctor_get(x_14, 7); -lean_inc(x_15); -lean_dec(x_14); -x_16 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_15, x_1); -x_17 = lean_box(x_16); -lean_ctor_set(x_11, 0, x_17); -return x_11; +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_11); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_20, 7); -lean_inc(x_21); -lean_dec(x_20); -x_22 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_21, x_1); -x_23 = lean_box(x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_19); -return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = lean_st_ref_take(x_8, x_11); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3; +x_12 = lean_st_mk_ref(x_11, x_10); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -x_14 = lean_ctor_get(x_13, 0); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); lean_dec(x_12); -x_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = !lean_is_exclusive(x_14); -if (x_18 == 0) +lean_inc(x_13); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_filterENodes___lambda__1___boxed), 12, 2); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_13); +x_16 = l_Lean_Meta_Grind_forEachENode(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_14, 7); -x_20 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_19, x_1, x_2); -lean_ctor_set(x_14, 7, x_20); -x_21 = lean_st_ref_set(x_8, x_13, x_15); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_st_ref_get(x_13, x_17); +lean_dec(x_13); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -x_24 = lean_box(0); -lean_ctor_set(x_21, 0, x_24); -return x_21; +return x_18; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_28 = lean_ctor_get(x_14, 0); -x_29 = lean_ctor_get(x_14, 1); -x_30 = lean_ctor_get(x_14, 2); -x_31 = lean_ctor_get(x_14, 3); -x_32 = lean_ctor_get(x_14, 4); -x_33 = lean_ctor_get(x_14, 5); -x_34 = lean_ctor_get(x_14, 6); -x_35 = lean_ctor_get(x_14, 7); -x_36 = lean_ctor_get(x_14, 8); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_14); -x_37 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_35, x_1, x_2); -x_38 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_38, 0, x_28); -lean_ctor_set(x_38, 1, x_29); -lean_ctor_set(x_38, 2, x_30); -lean_ctor_set(x_38, 3, x_31); -lean_ctor_set(x_38, 4, x_32); -lean_ctor_set(x_38, 5, x_33); -lean_ctor_set(x_38, 6, x_34); -lean_ctor_set(x_38, 7, x_37); -lean_ctor_set(x_38, 8, x_36); -lean_ctor_set(x_13, 0, x_38); -x_39 = lean_st_ref_set(x_8, x_13, x_15); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_41 = x_39; -} else { - lean_dec_ref(x_39); - x_41 = lean_box(0); -} -x_42 = lean_box(0); -if (lean_is_scalar(x_41)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_41; -} -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; -} +uint8_t x_23; +lean_dec(x_13); +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) +{ +return x_16; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_44 = lean_ctor_get(x_13, 1); -x_45 = lean_ctor_get(x_13, 2); -x_46 = lean_ctor_get(x_13, 3); -x_47 = lean_ctor_get(x_13, 4); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_13); -x_48 = lean_ctor_get(x_14, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_14, 1); -lean_inc(x_49); -x_50 = lean_ctor_get(x_14, 2); -lean_inc(x_50); -x_51 = lean_ctor_get(x_14, 3); -lean_inc(x_51); -x_52 = lean_ctor_get(x_14, 4); -lean_inc(x_52); -x_53 = lean_ctor_get(x_14, 5); -lean_inc(x_53); -x_54 = lean_ctor_get(x_14, 6); -lean_inc(x_54); -x_55 = lean_ctor_get(x_14, 7); -lean_inc(x_55); -x_56 = lean_ctor_get(x_14, 8); -lean_inc(x_56); -if (lean_is_exclusive(x_14)) { - lean_ctor_release(x_14, 0); - lean_ctor_release(x_14, 1); - lean_ctor_release(x_14, 2); - lean_ctor_release(x_14, 3); - lean_ctor_release(x_14, 4); - lean_ctor_release(x_14, 5); - lean_ctor_release(x_14, 6); - lean_ctor_release(x_14, 7); - lean_ctor_release(x_14, 8); - x_57 = x_14; -} else { - lean_dec_ref(x_14); - x_57 = lean_box(0); -} -x_58 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_55, x_1, x_2); -if (lean_is_scalar(x_57)) { - x_59 = lean_alloc_ctor(0, 9, 0); -} else { - x_59 = x_57; -} -lean_ctor_set(x_59, 0, x_48); -lean_ctor_set(x_59, 1, x_49); -lean_ctor_set(x_59, 2, x_50); -lean_ctor_set(x_59, 3, x_51); -lean_ctor_set(x_59, 4, x_52); -lean_ctor_set(x_59, 5, x_53); -lean_ctor_set(x_59, 6, x_54); -lean_ctor_set(x_59, 7, x_58); -lean_ctor_set(x_59, 8, x_56); -x_60 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_44); -lean_ctor_set(x_60, 2, x_45); -lean_ctor_set(x_60, 3, x_46); -lean_ctor_set(x_60, 4, x_47); -x_61 = lean_st_ref_set(x_8, x_60, x_15); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_63 = x_61; -} else { - lean_dec_ref(x_61); - x_63 = lean_box(0); -} -x_64 = lean_box(0); -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_63; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_16, 0); +x_25 = lean_ctor_get(x_16, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_16); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_62); -return x_65; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_11 = l_Lean_Meta_Grind_markAsInconsistent(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_st_ref_get(x_2, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_filterENodes___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_6, x_5); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_15); lean_dec(x_14); -x_17 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_7); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +else { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_dec(x_7); +x_19 = lean_array_uget(x_4, x_6); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_17); -lean_inc(x_16); -x_21 = l_Lean_MVarId_getType(x_16, x_6, x_7, x_8, x_9, x_20); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); +x_21 = lean_ctor_get(x_19, 2); +lean_inc(x_21); +x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_20, x_21); lean_dec(x_21); -lean_inc(x_22); -x_24 = l_Lean_Expr_isFalse(x_22); -if (x_24 == 0) +lean_dec(x_20); +if (x_22 == 0) { -lean_object* x_25; +size_t x_23; size_t x_24; lean_object* x_25; +lean_dec(x_19); +x_23 = 1; +x_24 = lean_usize_add(x_6, x_23); +x_25 = lean_box(0); +x_6 = x_24; +x_7 = x_25; +goto _start; +} +else +{ +lean_object* x_27; +lean_inc(x_1); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_25 = l_Lean_Meta_mkFalseElim(x_22, x_1, x_6, x_7, x_8, x_9, x_23); -if (lean_obj_tag(x_25) == 0) +x_27 = lean_apply_10(x_1, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_16, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_28; +lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = 1; +x_30 = lean_usize_add(x_6, x_29); +x_31 = lean_box(0); +x_6 = x_30; +x_7 = x_31; +x_16 = x_28; +goto _start; } else { -uint8_t x_29; -lean_dec(x_16); +uint8_t x_33; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_29 = !lean_is_exclusive(x_25); -if (x_29 == 0) +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_27); +if (x_33 == 0) { -return x_25; +return x_27; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_25, 0); -x_31 = lean_ctor_get(x_25, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_25); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_27, 0); +x_35 = lean_ctor_get(x_27, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_27); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachEqc(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; +x_11 = l_Lean_Meta_Grind_getENodes(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_box(0); +x_15 = lean_array_size(x_12); +x_16 = 0; +x_17 = lean_box(0); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(x_1, x_12, x_14, x_12, x_15, x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_12); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_17); +return x_18; +} else { -lean_object* x_33; -lean_dec(x_22); -x_33 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_16, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_33; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_17); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } else { -uint8_t x_34; -lean_dec(x_16); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_34 = !lean_is_exclusive(x_21); -if (x_34 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) { -return x_21; +return x_18; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_21, 0); -x_36 = lean_ctor_get(x_21, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_21); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_18, 0); +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_18); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: { -uint8_t x_38; -lean_dec(x_16); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_18 = lean_unbox_usize(x_6); lean_dec(x_6); -lean_dec(x_1); -x_38 = !lean_is_exclusive(x_17); -if (x_38 == 0) +x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg___boxed), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_17, 0); -lean_dec(x_39); -x_40 = lean_box(0); -lean_ctor_set(x_17, 0, x_40); -return x_17; +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; } -else +} +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__1() { +_start: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_17, 1); -lean_inc(x_41); -lean_dec(x_17); -x_42 = lean_box(0); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_41); -return x_43; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_instInhabitedMethods___lambda__1___boxed), 10, 0); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg___boxed), 10, 1); +lean_closure_set(x_2, 0, x_1); +x_3 = l_Lean_Meta_Grind_instInhabitedMethods___closed__1; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set(x_4, 2, x_2); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_MVarId_isAssigned___at_Lean_Meta_Grind_closeGoal___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_ReaderT_pure___at_Lean_Meta_Grind_instInhabitedMethods___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -17149,16 +21902,14 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; -x_12 = l_Lean_MVarId_assign___at_Lean_Meta_Grind_closeGoal___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_instInhabitedMethods___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -17166,329 +21917,435 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_12; +lean_dec(x_2); +lean_dec(x_1); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_closeGoal___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(lean_object* x_1) { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_closeGoal(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; +lean_inc(x_1); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1___boxed(lean_object* x_1) { _start: { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_1, x_3); -return x_4; +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(x_1); +lean_dec(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef(lean_object* x_1) { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -x_5 = lean_array_push(x_1, x_4); -return x_5; +lean_inc(x_1); +return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef___boxed(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); +lean_object* x_2; +x_2 = l_Lean_Meta_Grind_Methods_toMethodsRef(x_1); +lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___lambda__1), 3, 0); +lean_inc(x_1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1___boxed(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2; -x_3 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; -x_4 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_1, x_3); -return x_4; +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(x_1); +lean_dec(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(lean_object* x_1) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods___boxed(lean_object* x_1) { +_start: { -return x_3; +lean_object* x_2; +x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(x_1); +lean_dec(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_ctor_get(x_5, 1); -lean_inc(x_8); +lean_object* x_9; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_Grind_getMethods(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_3, 0); +lean_inc(x_11); +x_12 = lean_apply_10(x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +x_12 = lean_apply_10(x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_applyFallback(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +x_11 = lean_apply_9(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_2); +x_13 = l_Lean_Meta_Grind_getNext(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, x_3); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_15); +if (x_18 == 0) +{ +lean_free_object(x_13); +x_2 = x_15; +x_3 = x_17; +x_12 = x_16; +goto _start; +} +else +{ +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_17); +return x_13; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_ctor_get(x_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_2); +lean_ctor_set(x_22, 1, x_3); +x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_20); +if (x_23 == 0) +{ +x_2 = x_20; +x_3 = x_22; +x_12 = x_21; goto _start; } +else +{ +lean_object* x_25; +lean_dec(x_20); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_21); +return x_25; +} +} +} +else +{ +uint8_t x_26; +lean_dec(x_3); +lean_dec(x_2); +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) +{ +return x_13; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_13, 0); +x_28 = lean_ctor_get(x_13, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_13); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} } } -LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_ctor_get(x_1, 7); -x_4 = lean_ctor_get(x_2, 7); -x_5 = lean_nat_dec_lt(x_3, x_4); -return x_5; +lean_object* x_13; +x_13 = l_Lean_Meta_Grind_getEqc_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_13; } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed), 2, 0); -return x_1; +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +lean_inc(x_1); +x_12 = l_Lean_Meta_Grind_getEqc_go(x_1, x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_7; -x_7 = lean_nat_dec_lt(x_3, x_4); -if (x_7 == 0) -{ -lean_dec(x_3); -return x_2; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_8 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1; -lean_inc(x_3); -x_9 = l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(x_1, x_2, x_8, x_3, x_4, lean_box(0), lean_box(0)); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_getEqc(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); -x_12 = lean_nat_dec_le(x_4, x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_1, x_11, x_3, x_10, lean_box(0), lean_box(0)); -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_add(x_10, x_14); -lean_dec(x_10); -x_2 = x_13; -x_3 = x_15; -x_5 = lean_box(0); -x_6 = lean_box(0); -goto _start; -} -else -{ -lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); return x_11; } } -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_10; uint8_t x_11; -x_10 = lean_st_ref_get(x_1, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_13); -lean_dec(x_13); -x_15 = lean_array_size(x_14); -x_16 = 0; -x_17 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_15, x_16, x_14); -x_18 = lean_array_get_size(x_17); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_sub(x_18, x_19); -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_nat_dec_eq(x_18, x_21); -if (x_22 == 0) -{ -uint8_t x_23; -x_23 = lean_nat_dec_le(x_21, x_20); -if (x_23 == 0) +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_5, x_4); +if (x_16 == 0) { -lean_object* x_24; -lean_inc(x_20); -x_24 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_18, x_17, x_20, x_20, lean_box(0), lean_box(0)); -lean_dec(x_20); -lean_dec(x_18); -lean_ctor_set(x_10, 0, x_24); -return x_10; +lean_object* x_17; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_15); +return x_17; } else { -lean_object* x_25; -x_25 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_18, x_17, x_21, x_20, lean_box(0), lean_box(0)); -lean_dec(x_20); +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_array_uget(x_3, x_5); +x_19 = lean_ctor_get(x_18, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); lean_dec(x_18); -lean_ctor_set(x_10, 0, x_25); -return x_10; -} -} -else +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_19, x_20); +lean_dec(x_19); +if (x_21 == 0) { +size_t x_22; size_t x_23; lean_dec(x_20); -lean_dec(x_18); -lean_ctor_set(x_10, 0, x_17); -return x_10; -} +x_22 = 1; +x_23 = lean_usize_add(x_5, x_22); +x_5 = x_23; +goto _start; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_26 = lean_ctor_get(x_10, 0); -x_27 = lean_ctor_get(x_10, 1); -lean_inc(x_27); +lean_object* x_25; +x_25 = l_Lean_Meta_Grind_getEqc(x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; +x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); -lean_dec(x_10); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_28); -lean_dec(x_28); -x_30 = lean_array_size(x_29); -x_31 = 0; -x_32 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_30, x_31, x_29); -x_33 = lean_array_get_size(x_32); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_sub(x_33, x_34); -x_36 = lean_unsigned_to_nat(0u); -x_37 = lean_nat_dec_eq(x_33, x_36); -if (x_37 == 0) +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_6); +x_29 = 1; +x_30 = lean_usize_add(x_5, x_29); +x_5 = x_30; +x_6 = x_28; +x_15 = x_27; +goto _start; +} +else { -uint8_t x_38; -x_38 = lean_nat_dec_le(x_36, x_35); -if (x_38 == 0) +uint8_t x_32; +lean_dec(x_6); +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) { -lean_object* x_39; lean_object* x_40; -lean_inc(x_35); -x_39 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_33, x_32, x_35, x_35, lean_box(0), lean_box(0)); -lean_dec(x_35); -lean_dec(x_33); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_27); -return x_40; +return x_25; } else { -lean_object* x_41; lean_object* x_42; -x_41 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_33, x_32, x_36, x_35, lean_box(0), lean_box(0)); -lean_dec(x_35); -lean_dec(x_33); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_27); -return x_42; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_25, 0); +x_34 = lean_ctor_get(x_25, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_25); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} } } +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; +x_10 = lean_box(0); +x_11 = l_Lean_Meta_Grind_getENodes(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_box(0); +x_15 = lean_array_size(x_12); +x_16 = 0; +x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(x_12, x_14, x_12, x_15, x_16, x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +lean_dec(x_12); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_17; +} else { -lean_object* x_43; -lean_dec(x_35); -lean_dec(x_33); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_32); -lean_ctor_set(x_43, 1, x_27); -return x_43; -} -} +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Grind_getENodes___spec__2(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___boxed(lean_object* x_1) { -_start: +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) { -lean_object* x_2; -x_2 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1(x_1); -lean_dec(x_1); -return x_2; -} +return x_17; } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_getENodes___spec__3(x_4, x_5, x_3); -return x_6; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___lambda__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_7; -x_7 = l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4(x_1, x_2, x_3, x_4, x_5, x_6); +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_4); lean_dec(x_4); +x_17 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(x_1, x_2, x_3, x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getENodes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Meta_Grind_getENodes(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Meta_Grind_getEqcs(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -17500,1001 +22357,1093 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_17; -x_17 = lean_usize_dec_lt(x_6, x_5); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_7); -x_19 = lean_array_uget(x_4, x_6); -lean_inc(x_1); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_20 = lean_apply_10(x_1, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_20) == 0) +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) { -lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = 1; -x_23 = lean_usize_add(x_6, x_22); -x_24 = lean_box(0); -x_6 = x_23; -x_7 = x_24; -x_16 = x_21; -goto _start; +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; +return x_8; } else { -uint8_t x_26; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_5, x_9); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_1); -x_26 = !lean_is_exclusive(x_20); -if (x_26 == 0) +if (x_10 == 0) { -return x_20; +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; +goto _start; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_20, 0); -x_28 = lean_ctor_get(x_20, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_20); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachENode(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; -x_11 = l_Lean_Meta_Grind_getENodes(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = lean_array_size(x_12); -x_16 = 0; -x_17 = lean_box(0); -x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(x_1, x_12, x_14, x_12, x_15, x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -lean_dec(x_12); -if (lean_obj_tag(x_18) == 0) +if (lean_obj_tag(x_1) == 0) { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +switch (lean_obj_tag(x_10)) { +case 0: { -lean_object* x_20; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -lean_ctor_set(x_18, 0, x_17); -return x_18; +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_11); +lean_dec(x_11); +return x_12; } -else +case 1: { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_17); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_13; size_t x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_usize_shift_right(x_2, x_5); +x_1 = x_13; +x_2 = x_14; +goto _start; +} +default: +{ +uint8_t x_16; +x_16 = 0; +return x_16; } } -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_18); -if (x_23 == 0) -{ -return x_18; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_18, 0); -x_25 = lean_ctor_get(x_18, 1); -lean_inc(x_25); -lean_inc(x_24); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 1); +lean_inc(x_18); +lean_dec(x_1); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3(x_17, x_18, lean_box(0), x_19, x_3); lean_dec(x_18); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; -} +lean_dec(x_17); +return x_20; } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachENode___spec__1(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_19; +uint64_t x_3; size_t x_4; uint8_t x_5; +x_3 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_4 = lean_uint64_to_usize(x_3); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2(x_1, x_4, x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isResolvedCaseSplit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_13; -lean_inc(x_3); -x_13 = lean_apply_10(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_2, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_13, 0); +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 18); lean_inc(x_14); -x_15 = lean_unbox(x_14); -lean_dec(x_14); -if (x_15 == 0) -{ -uint8_t x_16; -lean_dec(x_3); -x_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_13, 0); -lean_dec(x_17); -x_18 = lean_box(0); -lean_ctor_set(x_13, 0, x_18); -return x_13; +lean_dec(x_13); +x_15 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(x_14, x_1); +x_16 = lean_box(x_15); +lean_ctor_set(x_11, 0, x_16); +return x_11; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 1); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_ctor_get(x_17, 18); lean_inc(x_19); -lean_dec(x_13); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; +lean_dec(x_17); +x_20 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(x_19, x_1); +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +return x_22; } } -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_dec(x_13); -x_23 = lean_st_ref_take(x_2, x_22); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_array_push(x_24, x_3); -x_27 = lean_st_ref_set(x_2, x_26, x_25); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -return x_27; +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_27, 0); -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_27); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +size_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__2(x_1, x_4, x_3); +lean_dec(x_3); +x_6 = lean_box(x_5); +return x_6; } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_Meta_Grind_isResolvedCaseSplit___spec__1(x_1, x_2); +lean_dec(x_2); +x_4 = lean_box(x_3); +return x_4; +} } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_isResolvedCaseSplit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_32; +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_isResolvedCaseSplit(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_32 = !lean_is_exclusive(x_13); -if (x_32 == 0) +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -return x_13; +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_dec(x_5); +return x_6; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_13, 0); -x_34 = lean_ctor_get(x_13, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_13); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} +lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_9); +x_12 = lean_uint64_to_usize(x_11); +x_13 = 1; +x_14 = lean_usize_sub(x_1, x_13); +x_15 = 5; +x_16 = lean_usize_mul(x_15, x_14); +x_17 = lean_usize_shift_right(x_12, x_16); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_5, x_18); +lean_dec(x_5); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(x_6, x_17, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_19; +x_6 = x_20; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_11 = l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1; -x_12 = lean_st_mk_ref(x_11, x_10); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -lean_inc(x_13); -x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_filterENodes___lambda__1___boxed), 12, 2); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_13); -x_16 = l_Lean_Meta_Grind_forEachENode(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_16) == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_st_ref_get(x_13, x_17); -lean_dec(x_13); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) { -return x_18; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } else { -uint8_t x_23; -lean_dec(x_13); -x_23 = !lean_is_exclusive(x_16); -if (x_23 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) { -return x_16; +lean_object* x_19; lean_object* x_20; +lean_dec(x_6); +lean_dec(x_5); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_2); +x_2 = x_20; +goto _start; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_16, 0); -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_16); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_1, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_1, 0); +lean_dec(x_24); +x_25 = lean_array_fset(x_5, x_2, x_3); +x_26 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); +return x_1; } +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_filterENodes___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_filterENodes___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_13; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_17; -x_17 = lean_usize_dec_lt(x_6, x_5); -if (x_17 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); lean_dec(x_13); +if (x_14 == 0) +{ lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_7); -lean_ctor_set(x_18, 1, x_16); -return x_18; +lean_dec(x_5); +lean_dec(x_4); +return x_1; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -lean_dec(x_7); -x_19 = lean_array_uget(x_4, x_6); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 2); -lean_inc(x_21); -x_22 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_20, x_21); -lean_dec(x_21); -lean_dec(x_20); -if (x_22 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -size_t x_23; size_t x_24; lean_object* x_25; -lean_dec(x_19); -x_23 = 1; -x_24 = lean_usize_add(x_6, x_23); -x_25 = lean_box(0); -x_6 = x_24; -x_7 = x_25; -goto _start; -} -else +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -lean_object* x_27; -lean_inc(x_1); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_27 = lean_apply_10(x_1, x_19, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_27) == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_19); +if (x_21 == 0) { -lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = 1; -x_30 = lean_usize_add(x_6, x_29); -x_31 = lean_box(0); -x_6 = x_30; -x_7 = x_31; -x_16 = x_28; -goto _start; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_array_fset(x_17, x_12, x_23); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_24); +return x_1; } else { -uint8_t x_33; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); +lean_object* x_25; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_25 = lean_array_fset(x_17, x_12, x_15); lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_1); -x_33 = !lean_is_exclusive(x_27); -if (x_33 == 0) -{ -return x_27; +lean_ctor_set(x_1, 0, x_25); +return x_1; +} } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_27, 0); -x_35 = lean_ctor_get(x_27, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_27); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_forEachEqc(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_26); +if (x_28 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; -x_11 = l_Lean_Meta_Grind_getENodes(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = lean_array_size(x_12); -x_16 = 0; -x_17 = lean_box(0); -x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(x_1, x_12, x_14, x_12, x_15, x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, x_4, x_5); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = lean_array_fset(x_17, x_12, x_30); lean_dec(x_12); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +lean_ctor_set(x_1, 0, x_31); +return x_1; +} +else { -lean_object* x_20; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -lean_ctor_set(x_18, 0, x_17); -return x_18; +lean_object* x_32; lean_object* x_33; +lean_dec(x_27); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_5); +x_33 = lean_array_fset(x_17, x_12, x_32); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_33); +return x_1; } -else -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_17); -lean_ctor_set(x_22, 1, x_21); -return x_22; } } -else +case 1: { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_18); -if (x_23 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -return x_18; +lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_usize_shift_right(x_2, x_9); +x_37 = lean_usize_add(x_3, x_8); +x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(x_35, x_36, x_37, x_4, x_5); +lean_ctor_set(x_15, 0, x_38); +x_39 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); +return x_1; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_18, 0); -x_25 = lean_ctor_get(x_18, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_18); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_15, 0); +lean_inc(x_40); +lean_dec(x_15); +x_41 = lean_usize_shift_right(x_2, x_9); +x_42 = lean_usize_add(x_3, x_8); +x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(x_40, x_41, x_42, x_4, x_5); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_array_fset(x_17, x_12, x_44); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_45); +return x_1; +} +} +default: +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_4); +lean_ctor_set(x_46, 1, x_5); +x_47 = lean_array_fset(x_17, x_12, x_46); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_47); +return x_1; } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: +else { -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_5); +lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +lean_dec(x_1); +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__2; +x_52 = lean_usize_land(x_2, x_51); +x_53 = lean_usize_to_nat(x_52); +x_54 = lean_array_get_size(x_48); +x_55 = lean_nat_dec_lt(x_53, x_54); +lean_dec(x_54); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_53); lean_dec(x_5); -x_18 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_19 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_forEachEqc___spec__1(x_1, x_2, x_3, x_4, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_19; -} +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -} -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__1() { -_start: +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_array_fget(x_48, x_53); +x_58 = lean_box(0); +x_59 = lean_array_fset(x_48, x_53, x_58); +switch (lean_obj_tag(x_57)) { +case 0: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_instInhabitedMethods___lambda__1___boxed), 10, 0); -return x_1; -} +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_62 = x_57; +} else { + lean_dec_ref(x_57); + x_62 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__2() { -_start: +x_63 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_4, x_60); +if (x_63 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Grind_instInhabitedMethods___closed__1; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_66 = lean_array_fset(x_59, x_53, x_65); +lean_dec(x_53); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; } -static lean_object* _init_l_Lean_Meta_Grind_instInhabitedMethods() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_Meta_Grind_instInhabitedMethods___closed__2; -return x_1; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_61); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_62; +} +lean_ctor_set(x_68, 0, x_4); +lean_ctor_set(x_68, 1, x_5); +x_69 = lean_array_fset(x_59, x_53, x_68); +lean_dec(x_53); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_instInhabitedMethods___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +case 1: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_instInhabitedMethods___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_11; +lean_object* x_71; lean_object* x_72; size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_71 = lean_ctor_get(x_57, 0); +lean_inc(x_71); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_72 = x_57; +} else { + lean_dec_ref(x_57); + x_72 = lean_box(0); } +x_73 = lean_usize_shift_right(x_2, x_50); +x_74 = lean_usize_add(x_3, x_49); +x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(x_71, x_73, x_74, x_4, x_5); +if (lean_is_scalar(x_72)) { + x_76 = lean_alloc_ctor(1, 1, 0); +} else { + x_76 = x_72; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(lean_object* x_1) { -_start: +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_array_fset(x_59, x_53, x_76); +lean_dec(x_53); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; +} +default: { -lean_inc(x_1); -return x_1; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_4); +lean_ctor_set(x_79, 1, x_5); +x_80 = lean_array_fset(x_59, x_53, x_79); +lean_dec(x_53); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_Methods_toMethodsRef_unsafe__1(x_1); -lean_dec(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef(lean_object* x_1) { -_start: -{ -lean_inc(x_1); -return x_1; } +else +{ +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; +x_83 = lean_unsigned_to_nat(0u); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__4(x_1, x_83, x_4, x_5); +x_85 = 7; +x_86 = lean_usize_dec_le(x_85, x_3); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_dec_lt(x_87, x_88); +lean_dec(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_84, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 1); +lean_inc(x_91); +lean_dec(x_84); +x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_Methods_toMethodsRef___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_Meta_Grind_Methods_toMethodsRef(x_1); -lean_dec(x_1); -return x_2; +return x_84; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(lean_object* x_1) { -_start: +else { -lean_inc(x_1); -return x_1; +return x_84; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods_unsafe__1(x_1); +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); +lean_inc(x_94); lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(lean_object* x_1) { -_start: +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_unsigned_to_nat(0u); +x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__4(x_96, x_97, x_4, x_5); +x_99 = 7; +x_100 = lean_usize_dec_le(x_99, x_3); +if (x_100 == 0) { -lean_inc(x_1); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods___boxed(lean_object* x_1) { -_start: +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_dec_lt(x_101, x_102); +lean_dec(x_101); +if (x_103 == 0) { -lean_object* x_2; -x_2 = l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_MethodsRef_toMethods(x_1); -lean_dec(x_1); -return x_2; -} +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_mkHCongrWithArity___spec__2___closed__3; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_8); -return x_9; +return x_98; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getMethods___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; -x_9 = l_Lean_Meta_Grind_getMethods(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; +return x_98; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_3, 0); -lean_inc(x_11); -x_12 = lean_apply_10(x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_propagateDown(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_3, 1); -lean_inc(x_11); -x_12 = lean_apply_10(x_11, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; +uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; +x_4 = l_Lean_Meta_Grind_instHashableENodeKey_unsafe__1(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = 1; +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(x_1, x_5, x_6, x_2, x_3); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; -lean_inc(x_2); -x_13 = l_Lean_Meta_Grind_getNext(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_2); -lean_ctor_set(x_17, 1, x_3); -x_18 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_15); -if (x_18 == 0) -{ -lean_free_object(x_13); -x_2 = x_15; -x_3 = x_17; -x_12 = x_16; -goto _start; -} -else -{ -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -} -else +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_take(x_3, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = lean_ctor_get(x_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_2); -lean_ctor_set(x_22, 1, x_3); -x_23 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_1, x_20); -if (x_23 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_13, 18); +x_17 = lean_box(0); +x_18 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__1(x_16, x_1, x_17); +lean_ctor_set(x_13, 18, x_18); +x_19 = lean_st_ref_set(x_3, x_13, x_14); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -x_2 = x_20; -x_3 = x_22; -x_12 = x_21; -goto _start; +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_17); +return x_19; } else { -lean_object* x_25; -lean_dec(x_20); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_22); -lean_ctor_set(x_25, 1, x_21); -return x_25; -} -} +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_17); +lean_ctor_set(x_23, 1, x_22); +return x_23; } -else -{ -uint8_t x_26; -lean_dec(x_3); -lean_dec(x_2); -x_26 = !lean_is_exclusive(x_13); -if (x_26 == 0) -{ -return x_13; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_13, 0); -x_28 = lean_ctor_get(x_13, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +x_26 = lean_ctor_get(x_13, 2); +x_27 = lean_ctor_get(x_13, 3); +x_28 = lean_ctor_get(x_13, 4); +x_29 = lean_ctor_get(x_13, 5); +x_30 = lean_ctor_get_uint8(x_13, sizeof(void*)*19); +x_31 = lean_ctor_get(x_13, 6); +x_32 = lean_ctor_get(x_13, 7); +x_33 = lean_ctor_get(x_13, 8); +x_34 = lean_ctor_get(x_13, 9); +x_35 = lean_ctor_get(x_13, 10); +x_36 = lean_ctor_get(x_13, 11); +x_37 = lean_ctor_get(x_13, 12); +x_38 = lean_ctor_get(x_13, 13); +x_39 = lean_ctor_get(x_13, 14); +x_40 = lean_ctor_get(x_13, 15); +x_41 = lean_ctor_get(x_13, 16); +x_42 = lean_ctor_get(x_13, 17); +x_43 = lean_ctor_get(x_13, 18); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); lean_dec(x_13); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +x_44 = lean_box(0); +x_45 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__1(x_43, x_1, x_44); +x_46 = lean_alloc_ctor(0, 19, 1); +lean_ctor_set(x_46, 0, x_24); +lean_ctor_set(x_46, 1, x_25); +lean_ctor_set(x_46, 2, x_26); +lean_ctor_set(x_46, 3, x_27); +lean_ctor_set(x_46, 4, x_28); +lean_ctor_set(x_46, 5, x_29); +lean_ctor_set(x_46, 6, x_31); +lean_ctor_set(x_46, 7, x_32); +lean_ctor_set(x_46, 8, x_33); +lean_ctor_set(x_46, 9, x_34); +lean_ctor_set(x_46, 10, x_35); +lean_ctor_set(x_46, 11, x_36); +lean_ctor_set(x_46, 12, x_37); +lean_ctor_set(x_46, 13, x_38); +lean_ctor_set(x_46, 14, x_39); +lean_ctor_set(x_46, 15, x_40); +lean_ctor_set(x_46, 16, x_41); +lean_ctor_set(x_46, 17, x_42); +lean_ctor_set(x_46, 18, x_45); +lean_ctor_set_uint8(x_46, sizeof(void*)*19, x_30); +x_47 = lean_st_ref_set(x_3, x_46, x_14); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); +} +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(0, 2, 0); +} else { + x_50 = x_49; } +lean_ctor_set(x_50, 0, x_44); +lean_ctor_set(x_50, 1, x_48); +return x_50; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +static lean_object* _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1() { _start: { -lean_object* x_13; -x_13 = l_Lean_Meta_Grind_getEqc_go(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("split", 5, 5); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2() { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -lean_inc(x_1); -x_12 = l_Lean_Meta_Grind_getEqc_go(x_1, x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_unchecked("resolved", 8, 8); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqc___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3() { _start: { -lean_object* x_11; -x_11 = l_Lean_Meta_Grind_getEqc(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_Grind_initFn____x40_Lean_Meta_Tactic_Grind_Types___hyg_78____closed__1; +x_2 = l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1; +x_3 = l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_5, x_4); -if (x_16 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = l_Lean_Meta_Grind_isResolvedCaseSplit(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) { -lean_object* x_17; -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_15); -return x_17; -} -else +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Grind_updateLastTag___spec__1(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_array_uget(x_3, x_5); -x_19 = lean_ctor_get(x_18, 2); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Meta_Grind_isSameExpr_unsafe__1(x_19, x_20); -lean_dec(x_19); -if (x_21 == 0) -{ -size_t x_22; size_t x_23; -lean_dec(x_20); -x_22 = 1; -x_23 = lean_usize_add(x_5, x_22); -x_5 = x_23; -goto _start; +lean_dec(x_16); +x_20 = lean_box(0); +x_21 = l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(x_1, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +return x_21; } else { -lean_object* x_25; -x_25 = l_Lean_Meta_Grind_getEqc(x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +uint8_t x_22; +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 1); +x_24 = lean_ctor_get(x_16, 0); +lean_dec(x_24); +x_25 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); if (lean_obj_tag(x_25) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; -x_26 = lean_ctor_get(x_25, 0); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_26 = lean_ctor_get(x_25, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); lean_dec(x_25); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_6); -x_29 = 1; -x_30 = lean_usize_add(x_5, x_29); -x_5 = x_30; -x_6 = x_28; -x_15 = x_27; -goto _start; +lean_inc(x_1); +x_27 = l_Lean_MessageData_ofExpr(x_1); +x_28 = l_Lean_Meta_Grind_getENode___closed__3; +lean_ctor_set_tag(x_16, 7); +lean_ctor_set(x_16, 1, x_27); +lean_ctor_set(x_16, 0, x_28); +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_16); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(x_1, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_32); +lean_dec(x_31); +return x_33; } else { -uint8_t x_32; -lean_dec(x_6); -x_32 = !lean_is_exclusive(x_25); -if (x_32 == 0) +uint8_t x_34; +lean_free_object(x_16); +lean_dec(x_1); +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) { return x_25; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_25, 0); -x_34 = lean_ctor_get(x_25, 1); -lean_inc(x_34); -lean_inc(x_33); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_ctor_get(x_25, 1); +lean_inc(x_36); +lean_inc(x_35); lean_dec(x_25); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -} +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; -x_10 = lean_box(0); -x_11 = l_Lean_Meta_Grind_getENodes(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = lean_array_size(x_12); -x_16 = 0; -x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(x_12, x_14, x_12, x_15, x_16, x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); -lean_dec(x_12); -if (lean_obj_tag(x_17) == 0) +else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_16, 1); +lean_inc(x_38); +lean_dec(x_16); +x_39 = l_Lean_Meta_Grind_updateLastTag(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +if (lean_obj_tag(x_39) == 0) { -return x_17; +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +lean_inc(x_1); +x_41 = l_Lean_MessageData_ofExpr(x_1); +x_42 = l_Lean_Meta_Grind_getENode___closed__3; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +x_45 = l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2(x_15, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(x_1, x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_47); +lean_dec(x_46); +return x_48; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_1); +x_49 = lean_ctor_get(x_39, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_39, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_51 = x_39; +} else { + lean_dec_ref(x_39); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} } } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) +uint8_t x_53; +lean_dec(x_1); +x_53 = !lean_is_exclusive(x_11); +if (x_53 == 0) { -return x_17; +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_11, 0); +lean_dec(x_54); +x_55 = lean_box(0); +lean_ctor_set(x_11, 0, x_55); +return x_11; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_17, 0); -x_24 = lean_ctor_get(x_17, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_17); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_11, 1); +lean_inc(x_56); +lean_dec(x_11); +x_57 = lean_box(0); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_56); +return x_58; } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_17 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_18 = l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_getEqcs___spec__1(x_1, x_2, x_3, x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__3(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Grind_markCaseSplitAsResolved___spec__2(x_1, x_6, x_7, x_4, x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Meta_Grind_markCaseSplitAsResolved___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_18; +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_getEqcs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_markCaseSplitAsResolved___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l_Lean_Meta_Grind_getEqcs(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_11; +x_11 = l_Lean_Meta_Grind_markCaseSplitAsResolved(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -18502,8 +23451,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_10; +return x_11; } } lean_object* initialize_Init_Grind_Tactics(uint8_t builtin, lean_object*); @@ -18628,114 +23576,116 @@ l_Lean_Meta_Grind_instInhabitedENode___closed__2 = _init_l_Lean_Meta_Grind_instI lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedENode___closed__2); l_Lean_Meta_Grind_instInhabitedENode = _init_l_Lean_Meta_Grind_instInhabitedENode(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedENode); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__1); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__2); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__3); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__4); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__5); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__6); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__7); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__8); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__9); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__10); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__11); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__12); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__13); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__14); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__15); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__16); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__17); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__18); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__19); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__20); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__21); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__22); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__23); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__24); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__25); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__26); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__27); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__28); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__29); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__30); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__31); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__32); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__33); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__34); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__35); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__36); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__37); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__38); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__39); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__40); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__41); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__42); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__43); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__44); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__45); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__46); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__47); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__48); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__49); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__50); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__51); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__52); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__53); -l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1532____closed__54); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__1); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__2); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__3); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__4); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__5); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__6); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__7); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__8); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__9); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__10); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__11); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__12); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__13); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__14); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__15); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__16); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__17); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__18); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__19); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__20); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__21); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__22); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__23); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__24); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__25); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__26); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__27); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__28); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__29); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__30); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__31); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__32); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__33); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__34); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__35); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__36); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__37); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__38); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__39); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__40); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__41); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__42); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__43); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__44); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__45); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__46); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__47); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__48); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__49); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__50); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__51); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__52); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__53); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__54); +l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55 = _init_l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_Grind_Types_0__Lean_Meta_Grind_reprENode____x40_Lean_Meta_Tactic_Grind_Types___hyg_1560____closed__55); l_Lean_Meta_Grind_instReprENode___closed__1 = _init_l_Lean_Meta_Grind_instReprENode___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instReprENode___closed__1); l_Lean_Meta_Grind_instReprENode = _init_l_Lean_Meta_Grind_instReprENode(); @@ -18744,6 +23694,10 @@ l_Lean_Meta_Grind_congrHash___closed__1 = _init_l_Lean_Meta_Grind_congrHash___cl lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__1); l_Lean_Meta_Grind_congrHash___closed__2 = _init_l_Lean_Meta_Grind_congrHash___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__2); +l_Lean_Meta_Grind_congrHash___closed__3 = _init_l_Lean_Meta_Grind_congrHash___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__3); +l_Lean_Meta_Grind_congrHash___closed__4 = _init_l_Lean_Meta_Grind_congrHash___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_congrHash___closed__4); l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Grind_instBEqPreInstance___spec__1___closed__1); l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_instBEqPreInstance___lambda__2___closed__1(); @@ -18758,6 +23712,8 @@ l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1 lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__1); l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2(); lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__2); +l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3 = _init_l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_Meta_Grind_instInhabitedGoal___spec__3); l_Lean_Meta_Grind_instInhabitedGoal___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__1); l_Lean_Meta_Grind_instInhabitedGoal___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__2(); @@ -18767,31 +23723,256 @@ l_Lean_Meta_Grind_instInhabitedGoal___closed__4 = _init_l_Lean_Meta_Grind_instIn lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__4); l_Lean_Meta_Grind_instInhabitedGoal___closed__5 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__5(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__5); +l_Lean_Meta_Grind_instInhabitedGoal___closed__6 = _init_l_Lean_Meta_Grind_instInhabitedGoal___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal___closed__6); l_Lean_Meta_Grind_instInhabitedGoal = _init_l_Lean_Meta_Grind_instInhabitedGoal(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedGoal); l_Lean_Meta_Grind_GoalM_run_x27___closed__1 = _init_l_Lean_Meta_Grind_GoalM_run_x27___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_GoalM_run_x27___closed__1); +l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__1(); +l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__2); +l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3 = _init_l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3(); +lean_mark_persistent(l_Lean_addTrace___at_Lean_Meta_Grind_updateLastTag___spec__2___closed__3); +l_Lean_Meta_Grind_updateLastTag___closed__1 = _init_l_Lean_Meta_Grind_updateLastTag___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__1); +l_Lean_Meta_Grind_updateLastTag___closed__2 = _init_l_Lean_Meta_Grind_updateLastTag___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__2); +l_Lean_Meta_Grind_updateLastTag___closed__3 = _init_l_Lean_Meta_Grind_updateLastTag___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__3); +l_Lean_Meta_Grind_updateLastTag___closed__4 = _init_l_Lean_Meta_Grind_updateLastTag___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__4); +l_Lean_Meta_Grind_updateLastTag___closed__5 = _init_l_Lean_Meta_Grind_updateLastTag___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_updateLastTag___closed__5); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__1); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__2); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__3); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__4); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__5); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__6); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__7); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__8); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__9); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__10); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__11); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__12); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__13); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__14); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__15); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__16); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__17); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__18); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__19); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__20); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__21); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__22); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__23); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24 = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d_______closed__24); +l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____ = _init_l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____(); +lean_mark_persistent(l_Lean_Meta_Grind_doElemTrace__goal_x5b___x5d____); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__1); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__2); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__3); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__4); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__5); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__6); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__7); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__8); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__9); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__10); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__11); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__12); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__13); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__14); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__15); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__16); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__17); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__18); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__19); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__20); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__21); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__22); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__23); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__24); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__25); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__26); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__27); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__28); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__29); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__30); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__31); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__32); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__33); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__34); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__35); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__36); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__37); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__38); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__39); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__40); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__41); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__42); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__43); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__44); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__45); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__46); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__47); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__48); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__49); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__50); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__51); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__52); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__53); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__54); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__55); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__56); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__57); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__58); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__59); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___lambda__1___closed__60); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__1); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__2); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__3); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__4); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__5); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__6); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__7); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__8); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__9); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__10); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__11); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__12); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__13); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__14); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__15); +l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16 = _init_l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16(); +lean_mark_persistent(l_Lean_Meta_Grind___aux__Lean__Meta__Tactic__Grind__Types______macroRules__Lean__Meta__Grind__doElemTrace__goal_x5b___x5d______1___closed__16); l_Lean_Meta_Grind_getENode___closed__1 = _init_l_Lean_Meta_Grind_getENode___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__1); l_Lean_Meta_Grind_getENode___closed__2 = _init_l_Lean_Meta_Grind_getENode___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__2); l_Lean_Meta_Grind_getENode___closed__3 = _init_l_Lean_Meta_Grind_getENode___closed__3(); lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__3); -l_Lean_Meta_Grind_getENode___closed__4 = _init_l_Lean_Meta_Grind_getENode___closed__4(); -lean_mark_persistent(l_Lean_Meta_Grind_getENode___closed__4); l_Lean_Meta_Grind_hasSameType___closed__1 = _init_l_Lean_Meta_Grind_hasSameType___closed__1(); +l_Lean_Meta_Grind_markAsInconsistent___closed__1 = _init_l_Lean_Meta_Grind_markAsInconsistent___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsInconsistent___closed__1); +l_Lean_Meta_Grind_markAsInconsistent___closed__2 = _init_l_Lean_Meta_Grind_markAsInconsistent___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsInconsistent___closed__2); +l_Lean_Meta_Grind_markAsInconsistent___closed__3 = _init_l_Lean_Meta_Grind_markAsInconsistent___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markAsInconsistent___closed__3); l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1 = _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1(); lean_mark_persistent(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__1); -l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2 = _init_l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2(); -lean_mark_persistent(l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Grind_getENodes___spec__1___closed__2); l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1 = _init_l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1(); lean_mark_persistent(l_Array_qsort_sort___at_Lean_Meta_Grind_getENodes___spec__4___closed__1); l_Lean_Meta_Grind_instInhabitedMethods___closed__1 = _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedMethods___closed__1); -l_Lean_Meta_Grind_instInhabitedMethods___closed__2 = _init_l_Lean_Meta_Grind_instInhabitedMethods___closed__2(); -lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedMethods___closed__2); l_Lean_Meta_Grind_instInhabitedMethods = _init_l_Lean_Meta_Grind_instInhabitedMethods(); lean_mark_persistent(l_Lean_Meta_Grind_instInhabitedMethods); +l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1 = _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__1); +l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2 = _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__2); +l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3 = _init_l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_markCaseSplitAsResolved___closed__3); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c b/stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c index 4576028b255a..81c560c88a86 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Grind/Util.c @@ -22,6 +22,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_byContra_x3f___lambda__1(lean_object*, le lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_clearAuxDecls___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7; static lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1___closed__3; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Grind_normalizeLevels___spec__1(lean_object*, lean_object*); @@ -42,12 +43,15 @@ static lean_object* l_Lean_Meta_Grind_foldProjs___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_unfoldReducible___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); static lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___closed__2; +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Meta_Grind_unfoldReducible___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_abstractNestedProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); @@ -60,10 +64,12 @@ static lean_object* l_Lean_MVarId_ensureNoMVar___closed__1; lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_ensureProp___closed__4; static lean_object* l_Lean_MVarId_unfoldReducible___closed__1; +lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_clearAuxDecls___spec__6___closed__4; static lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1___closed__4; lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_visit___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_byContra_x3f___lambda__1___closed__5; +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9; LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Meta_Grind_unfoldReducible___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_byContra_x3f___lambda__1___closed__4; lean_object* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -80,11 +86,13 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducib static lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_clearAuxDecls___spec__6___closed__1; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_eraseIrrelevantMData___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_clearAuxDecls___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_betaReduce___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_normalize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducible___spec__7___rarg___closed__3; static lean_object* l_Lean_MVarId_byContra_x3f___closed__1; static lean_object* l_Lean_MVarId_ensureProp___closed__1; @@ -101,6 +109,7 @@ uint8_t l_List_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Grind_normalizeLevels___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_betaReduce___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; static lean_object* l_Lean_MVarId_ensureNoMVar___closed__4; static lean_object* l_Lean_Core_transform___at_Lean_Meta_Grind_unfoldReducible___spec__1___closed__1; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducible___spec__7___rarg___closed__6; @@ -110,7 +119,9 @@ extern lean_object* l_Lean_levelZero; static lean_object* l_Lean_MVarId_byContra_x3f___closed__2; static lean_object* l_Lean_MVarId_ensureNoMVar___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_betaReduce(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; uint64_t l_Lean_Expr_hash(lean_object*); +lean_object* lean_grind_normalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_visit___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Core_betaReduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -128,6 +139,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_byContra_x3f___lambda__2(lean_object*, le LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_clearAuxDecls___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_isAuxDecl(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Grind_unfoldReducible___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Meta_Grind_unfoldReducible___spec__2___lambda__2(size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_ensureProp___closed__2; static lean_object* l_Lean_MVarId_ensureNoMVar___closed__6; @@ -149,6 +161,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_clearAuxDecls___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_clearAuxDecls___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform_visit_visitPost___at_Lean_Meta_Grind_unfoldReducible___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -165,6 +178,7 @@ static lean_object* l_Lean_MVarId_ensureNoMVar___closed__5; LEAN_EXPORT lean_object* l_Lean_MVarId_transformTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_MVarId_clearAuxDecls___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_visit___spec__2(lean_object*, lean_object*); @@ -177,14 +191,18 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_byContra_x3f___lambda__1___boxed(lean_obj LEAN_EXPORT lean_object* l_Lean_MVarId_ensureProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_clearAuxDecls___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5; +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_clearAuxDecls___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_array_size(lean_object*); static lean_object* l_Lean_MVarId_ensureProp___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Grind_unfoldReducible___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Grind_unfoldReducible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducible___spec__7(lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* l_ReaderT_bind___at_Lean_Meta_zetaReduce___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -206,9 +224,12 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_betaReduce___lambda__1___boxed(lean_objec static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Grind_unfoldReducible___spec__7___rarg___closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_MVarId_clearAuxDecls___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofName(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Meta_Grind_unfoldReducible___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_MVarId_clearAuxDecls___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4; static lean_object* l_Lean_Meta_Grind_unfoldReducible___closed__2; size_t lean_usize_land(size_t, size_t); lean_object* l_Lean_Core_betaReduce___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -217,6 +238,7 @@ LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Meta_Grind_unfold uint8_t l_Lean_Expr_hasExprMVar(lean_object*); static lean_object* l_Lean_Meta_Grind_normalizeLevels___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_clearAuxDecls___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_406_(uint8_t, uint8_t); static lean_object* _init_l_Lean_MVarId_ensureNoMVar___closed__1() { _start: @@ -5087,7 +5109,7 @@ lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); return x_6; } -case 7: +case 8: { lean_object* x_7; lean_object* x_8; x_7 = lean_alloc_ctor(0, 1, 0); @@ -5097,42 +5119,32 @@ lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_4); return x_8; } -case 8: -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_9, 0, x_1); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_4); -return x_10; -} case 10: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); lean_dec(x_1); -x_12 = lean_alloc_ctor(1, 1, 0); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_11, 0, x_10); +x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); -x_13 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_4); -return x_14; +lean_ctor_set(x_12, 1, x_4); +return x_12; } default: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_1); -x_16 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_16, 0, x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_4); -return x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_1); +x_14 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_14, 0, x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_4); +return x_15; } } } @@ -5195,7 +5207,122 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("issues", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_MVarId_ensureNoMVar___closed__1; +x_2 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found `Expr.proj` but `", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` is not marked as structure", 28, 28); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("found `Expr.proj` with invalid field index `", 44, 44); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { if (lean_obj_tag(x_1) == 11) @@ -5217,198 +5344,463 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); +lean_inc(x_7); x_15 = l_Lean_getStructureInfo_x3f(x_14, x_7); lean_dec(x_14); if (lean_obj_tag(x_15) == 0) { -lean_object* x_16; +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_dec(x_9); lean_dec(x_8); +x_16 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_17 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_16, x_2, x_3, x_4, x_5, x_13); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_10); +lean_dec(x_7); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_box(0); +x_22 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_21, x_2, x_3, x_4, x_5, x_20); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_16 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_16, 0, x_1); -lean_ctor_set(x_10, 0, x_16); -return x_10; +return x_22; } else { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_15, 0); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_array_get_size(x_19); -x_21 = lean_nat_dec_lt(x_8, x_20); -lean_dec(x_20); -if (x_21 == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_24 = lean_ctor_get(x_17, 1); +x_25 = lean_ctor_get(x_17, 0); +lean_dec(x_25); +x_26 = l_Lean_MessageData_ofName(x_7); +x_27 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4; +lean_ctor_set_tag(x_17, 7); +lean_ctor_set(x_17, 1, x_26); +lean_ctor_set(x_17, 0, x_27); +x_28 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_28); +lean_ctor_set(x_10, 0, x_17); +lean_inc(x_1); +x_29 = l_Lean_indentExpr(x_1); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_10); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_16, x_32, x_2, x_3, x_4, x_5, x_24); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_34, x_2, x_3, x_4, x_5, x_35); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_34); +return x_36; +} +else { -lean_dec(x_19); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_37 = lean_ctor_get(x_17, 1); +lean_inc(x_37); +lean_dec(x_17); +x_38 = l_Lean_MessageData_ofName(x_7); +x_39 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4; +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_41); +lean_ctor_set(x_10, 0, x_40); +lean_inc(x_1); +x_42 = l_Lean_indentExpr(x_1); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_10); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_16, x_45, x_2, x_3, x_4, x_5, x_37); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_47, x_2, x_3, x_4, x_5, x_48); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_47); +return x_49; +} +} +} +else +{ +uint8_t x_50; +lean_dec(x_7); +x_50 = !lean_is_exclusive(x_15); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_ctor_get(x_15, 0); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = lean_array_get_size(x_52); +x_54 = lean_nat_dec_lt(x_8, x_53); +lean_dec(x_53); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +lean_dec(x_52); lean_dec(x_9); +x_55 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_56 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_55, x_2, x_3, x_4, x_5, x_13); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_unbox(x_57); +lean_dec(x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_free_object(x_15); +lean_free_object(x_10); lean_dec(x_8); +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_box(0); +x_61 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_60, x_2, x_3, x_4, x_5, x_59); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_ctor_set_tag(x_15, 0); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_10, 0, x_15); -return x_10; +return x_61; } else { -lean_object* x_22; lean_object* x_23; +uint8_t x_62; +x_62 = !lean_is_exclusive(x_56); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_63 = lean_ctor_get(x_56, 1); +x_64 = lean_ctor_get(x_56, 0); +lean_dec(x_64); +x_65 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +lean_ctor_set_tag(x_15, 3); +lean_ctor_set(x_15, 0, x_65); +x_66 = l_Lean_MessageData_ofFormat(x_15); +x_67 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; +lean_ctor_set_tag(x_56, 7); +lean_ctor_set(x_56, 1, x_66); +lean_ctor_set(x_56, 0, x_67); +x_68 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_68); +lean_ctor_set(x_10, 0, x_56); +lean_inc(x_1); +x_69 = l_Lean_indentExpr(x_1); +x_70 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_70, 0, x_10); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_55, x_72, x_2, x_3, x_4, x_5, x_63); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_74, x_2, x_3, x_4, x_5, x_75); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_74); +return x_76; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_77 = lean_ctor_get(x_56, 1); +lean_inc(x_77); +lean_dec(x_56); +x_78 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +lean_ctor_set_tag(x_15, 3); +lean_ctor_set(x_15, 0, x_78); +x_79 = l_Lean_MessageData_ofFormat(x_15); +x_80 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; +x_81 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_79); +x_82 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_82); +lean_ctor_set(x_10, 0, x_81); +lean_inc(x_1); +x_83 = l_Lean_indentExpr(x_1); +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_10); +lean_ctor_set(x_84, 1, x_83); +x_85 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_86 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_55, x_86, x_2, x_3, x_4, x_5, x_77); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +x_90 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_88, x_2, x_3, x_4, x_5, x_89); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_88); +return x_90; +} +} +} +else +{ +lean_object* x_91; lean_object* x_92; lean_free_object(x_10); lean_dec(x_1); -x_22 = lean_array_fget(x_19, x_8); +x_91 = lean_array_fget(x_52, x_8); lean_dec(x_8); -lean_dec(x_19); -x_23 = l_Lean_Meta_mkProjection(x_9, x_22, x_2, x_3, x_4, x_5, x_13); -if (lean_obj_tag(x_23) == 0) +lean_dec(x_52); +x_92 = l_Lean_Meta_mkProjection(x_9, x_91, x_2, x_3, x_4, x_5, x_13); +if (lean_obj_tag(x_92) == 0) { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +uint8_t x_93; +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) { -lean_object* x_25; -x_25 = lean_ctor_get(x_23, 0); +lean_object* x_94; +x_94 = lean_ctor_get(x_92, 0); lean_ctor_set_tag(x_15, 0); -lean_ctor_set(x_15, 0, x_25); -lean_ctor_set(x_23, 0, x_15); -return x_23; +lean_ctor_set(x_15, 0, x_94); +lean_ctor_set(x_92, 0, x_15); +return x_92; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_23, 0); -x_27 = lean_ctor_get(x_23, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_23); +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_92, 0); +x_96 = lean_ctor_get(x_92, 1); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_92); lean_ctor_set_tag(x_15, 0); -lean_ctor_set(x_15, 0, x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_15); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_ctor_set(x_15, 0, x_95); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_15); +lean_ctor_set(x_97, 1, x_96); +return x_97; } } else { -uint8_t x_29; +uint8_t x_98; lean_free_object(x_15); -x_29 = !lean_is_exclusive(x_23); -if (x_29 == 0) +x_98 = !lean_is_exclusive(x_92); +if (x_98 == 0) { -return x_23; +return x_92; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_23, 0); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_23); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_92, 0); +x_100 = lean_ctor_get(x_92, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_92); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; } } } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_33 = lean_ctor_get(x_15, 0); -lean_inc(x_33); +lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_102 = lean_ctor_get(x_15, 0); +lean_inc(x_102); lean_dec(x_15); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = lean_array_get_size(x_34); -x_36 = lean_nat_dec_lt(x_8, x_35); -lean_dec(x_35); -if (x_36 == 0) -{ -lean_object* x_37; -lean_dec(x_34); +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +lean_dec(x_102); +x_104 = lean_array_get_size(x_103); +x_105 = lean_nat_dec_lt(x_8, x_104); +lean_dec(x_104); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +lean_dec(x_103); lean_dec(x_9); +x_106 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_107 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_106, x_2, x_3, x_4, x_5, x_13); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_unbox(x_108); +lean_dec(x_108); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_free_object(x_10); lean_dec(x_8); +x_110 = lean_ctor_get(x_107, 1); +lean_inc(x_110); +lean_dec(x_107); +x_111 = lean_box(0); +x_112 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_111, x_2, x_3, x_4, x_5, x_110); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_1); -lean_ctor_set(x_10, 0, x_37); -return x_10; +return x_112; } else { -lean_object* x_38; lean_object* x_39; +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_113 = lean_ctor_get(x_107, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_114 = x_107; +} else { + lean_dec_ref(x_107); + x_114 = lean_box(0); +} +x_115 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +x_116 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_116, 0, x_115); +x_117 = l_Lean_MessageData_ofFormat(x_116); +x_118 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; +if (lean_is_scalar(x_114)) { + x_119 = lean_alloc_ctor(7, 2, 0); +} else { + x_119 = x_114; + lean_ctor_set_tag(x_119, 7); +} +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_117); +x_120 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; +lean_ctor_set_tag(x_10, 7); +lean_ctor_set(x_10, 1, x_120); +lean_ctor_set(x_10, 0, x_119); +lean_inc(x_1); +x_121 = l_Lean_indentExpr(x_1); +x_122 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_122, 0, x_10); +lean_ctor_set(x_122, 1, x_121); +x_123 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_124 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +x_125 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_106, x_124, x_2, x_3, x_4, x_5, x_113); +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec(x_125); +x_128 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_126, x_2, x_3, x_4, x_5, x_127); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_126); +return x_128; +} +} +else +{ +lean_object* x_129; lean_object* x_130; lean_free_object(x_10); lean_dec(x_1); -x_38 = lean_array_fget(x_34, x_8); +x_129 = lean_array_fget(x_103, x_8); lean_dec(x_8); -lean_dec(x_34); -x_39 = l_Lean_Meta_mkProjection(x_9, x_38, x_2, x_3, x_4, x_5, x_13); -if (lean_obj_tag(x_39) == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_42 = x_39; +lean_dec(x_103); +x_130 = l_Lean_Meta_mkProjection(x_9, x_129, x_2, x_3, x_4, x_5, x_13); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_133 = x_130; } else { - lean_dec_ref(x_39); - x_42 = lean_box(0); + lean_dec_ref(x_130); + x_133 = lean_box(0); } -x_43 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_43, 0, x_40); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); +x_134 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_134, 0, x_131); +if (lean_is_scalar(x_133)) { + x_135 = lean_alloc_ctor(0, 2, 0); } else { - x_44 = x_42; + x_135 = x_133; } -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_41); -return x_44; +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_132); +return x_135; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_39, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_39, 1); -lean_inc(x_46); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_47 = x_39; +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_136 = lean_ctor_get(x_130, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_130, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_138 = x_130; } else { - lean_dec_ref(x_39); - x_47 = lean_box(0); + lean_dec_ref(x_130); + x_138 = lean_box(0); } -if (lean_is_scalar(x_47)) { - x_48 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_138)) { + x_139 = lean_alloc_ctor(1, 2, 0); } else { - x_48 = x_47; + x_139 = x_138; } -lean_ctor_set(x_48, 0, x_45); -lean_ctor_set(x_48, 1, x_46); -return x_48; +lean_ctor_set(x_139, 0, x_136); +lean_ctor_set(x_139, 1, x_137); +return x_139; } } } @@ -5416,136 +5808,263 @@ return x_48; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_49 = lean_ctor_get(x_10, 0); -x_50 = lean_ctor_get(x_10, 1); -lean_inc(x_50); -lean_inc(x_49); +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_140 = lean_ctor_get(x_10, 0); +x_141 = lean_ctor_get(x_10, 1); +lean_inc(x_141); +lean_inc(x_140); lean_dec(x_10); -x_51 = lean_ctor_get(x_49, 0); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_getStructureInfo_x3f(x_51, x_7); -lean_dec(x_51); -if (lean_obj_tag(x_52) == 0) +x_142 = lean_ctor_get(x_140, 0); +lean_inc(x_142); +lean_dec(x_140); +lean_inc(x_7); +x_143 = l_Lean_getStructureInfo_x3f(x_142, x_7); +lean_dec(x_142); +if (lean_obj_tag(x_143) == 0) { -lean_object* x_53; lean_object* x_54; +lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; lean_dec(x_9); lean_dec(x_8); +x_144 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_145 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_144, x_2, x_3, x_4, x_5, x_141); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_unbox(x_146); +lean_dec(x_146); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_7); +x_148 = lean_ctor_get(x_145, 1); +lean_inc(x_148); +lean_dec(x_145); +x_149 = lean_box(0); +x_150 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_149, x_2, x_3, x_4, x_5, x_148); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_53 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_53, 0, x_1); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_50); -return x_54; +return x_150; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_55 = lean_ctor_get(x_52, 0); -lean_inc(x_55); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - x_56 = x_52; +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_151 = lean_ctor_get(x_145, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_152 = x_145; } else { - lean_dec_ref(x_52); - x_56 = lean_box(0); + lean_dec_ref(x_145); + x_152 = lean_box(0); } -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_array_get_size(x_57); -x_59 = lean_nat_dec_lt(x_8, x_58); -lean_dec(x_58); -if (x_59 == 0) +x_153 = l_Lean_MessageData_ofName(x_7); +x_154 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4; +if (lean_is_scalar(x_152)) { + x_155 = lean_alloc_ctor(7, 2, 0); +} else { + x_155 = x_152; + lean_ctor_set_tag(x_155, 7); +} +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_153); +x_156 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6; +x_157 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +lean_inc(x_1); +x_158 = l_Lean_indentExpr(x_1); +x_159 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_161 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +x_162 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_144, x_161, x_2, x_3, x_4, x_5, x_151); +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_163, x_2, x_3, x_4, x_5, x_164); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_163); +return x_165; +} +} +else { -lean_object* x_60; lean_object* x_61; -lean_dec(x_57); +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; uint8_t x_170; +lean_dec(x_7); +x_166 = lean_ctor_get(x_143, 0); +lean_inc(x_166); +if (lean_is_exclusive(x_143)) { + lean_ctor_release(x_143, 0); + x_167 = x_143; +} else { + lean_dec_ref(x_143); + x_167 = lean_box(0); +} +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec(x_166); +x_169 = lean_array_get_size(x_168); +x_170 = lean_nat_dec_lt(x_8, x_169); +lean_dec(x_169); +if (x_170 == 0) +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; +lean_dec(x_168); lean_dec(x_9); +x_171 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2; +x_172 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_171, x_2, x_3, x_4, x_5, x_141); +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_unbox(x_173); +lean_dec(x_173); +if (x_174 == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_dec(x_167); lean_dec(x_8); +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +lean_dec(x_172); +x_176 = lean_box(0); +x_177 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_176, x_2, x_3, x_4, x_5, x_175); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -if (lean_is_scalar(x_56)) { - x_60 = lean_alloc_ctor(0, 1, 0); +return x_177; +} +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_178 = lean_ctor_get(x_172, 1); +lean_inc(x_178); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + x_179 = x_172; } else { - x_60 = x_56; - lean_ctor_set_tag(x_60, 0); + lean_dec_ref(x_172); + x_179 = lean_box(0); +} +x_180 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +if (lean_is_scalar(x_167)) { + x_181 = lean_alloc_ctor(3, 1, 0); +} else { + x_181 = x_167; + lean_ctor_set_tag(x_181, 3); +} +lean_ctor_set(x_181, 0, x_180); +x_182 = l_Lean_MessageData_ofFormat(x_181); +x_183 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10; +if (lean_is_scalar(x_179)) { + x_184 = lean_alloc_ctor(7, 2, 0); +} else { + x_184 = x_179; + lean_ctor_set_tag(x_184, 7); +} +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_182); +x_185 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12; +x_186 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +lean_inc(x_1); +x_187 = l_Lean_indentExpr(x_1); +x_188 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set(x_188, 1, x_187); +x_189 = l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8; +x_190 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +x_191 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_171, x_190, x_2, x_3, x_4, x_5, x_178); +x_192 = lean_ctor_get(x_191, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_191, 1); +lean_inc(x_193); +lean_dec(x_191); +x_194 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_192, x_2, x_3, x_4, x_5, x_193); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_192); +return x_194; } -lean_ctor_set(x_60, 0, x_1); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_50); -return x_61; } else { -lean_object* x_62; lean_object* x_63; +lean_object* x_195; lean_object* x_196; lean_dec(x_1); -x_62 = lean_array_fget(x_57, x_8); +x_195 = lean_array_fget(x_168, x_8); lean_dec(x_8); -lean_dec(x_57); -x_63 = l_Lean_Meta_mkProjection(x_9, x_62, x_2, x_3, x_4, x_5, x_50); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_66 = x_63; +lean_dec(x_168); +x_196 = l_Lean_Meta_mkProjection(x_9, x_195, x_2, x_3, x_4, x_5, x_141); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +if (lean_is_exclusive(x_196)) { + lean_ctor_release(x_196, 0); + lean_ctor_release(x_196, 1); + x_199 = x_196; } else { - lean_dec_ref(x_63); - x_66 = lean_box(0); + lean_dec_ref(x_196); + x_199 = lean_box(0); } -if (lean_is_scalar(x_56)) { - x_67 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_167)) { + x_200 = lean_alloc_ctor(0, 1, 0); } else { - x_67 = x_56; - lean_ctor_set_tag(x_67, 0); + x_200 = x_167; + lean_ctor_set_tag(x_200, 0); } -lean_ctor_set(x_67, 0, x_64); -if (lean_is_scalar(x_66)) { - x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_197); +if (lean_is_scalar(x_199)) { + x_201 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_66; + x_201 = x_199; } -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_65); -return x_68; +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_198); +return x_201; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -lean_dec(x_56); -x_69 = lean_ctor_get(x_63, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_63, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_71 = x_63; +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_167); +x_202 = lean_ctor_get(x_196, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_196, 1); +lean_inc(x_203); +if (lean_is_exclusive(x_196)) { + lean_ctor_release(x_196, 0); + lean_ctor_release(x_196, 1); + x_204 = x_196; } else { - lean_dec_ref(x_63); - x_71 = lean_box(0); + lean_dec_ref(x_196); + x_204 = lean_box(0); } -if (lean_is_scalar(x_71)) { - x_72 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_204)) { + x_205 = lean_alloc_ctor(1, 2, 0); } else { - x_72 = x_71; + x_205 = x_204; } -lean_ctor_set(x_72, 0, x_69); -lean_ctor_set(x_72, 1, x_70); -return x_72; +lean_ctor_set(x_205, 0, x_202); +lean_ctor_set(x_205, 1, x_203); +return x_205; } } } @@ -5553,21 +6072,21 @@ return x_72; } else { -lean_object* x_73; lean_object* x_74; +lean_object* x_206; lean_object* x_207; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_73 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_73, 0, x_1); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_6); -return x_74; +x_206 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_206, 0, x_1); +x_207 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_207, 0, x_206); +lean_ctor_set(x_207, 1, x_6); +return x_207; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; @@ -5582,7 +6101,7 @@ static lean_object* _init_l_Lean_Meta_Grind_foldProjs___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__3___boxed), 6, 0); return x_1; } } @@ -5590,7 +6109,7 @@ static lean_object* _init_l_Lean_Meta_Grind_foldProjs___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__1), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Grind_foldProjs___lambda__2), 6, 0); return x_1; } } @@ -5605,11 +6124,24 @@ x_10 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_1, x_7, x_8, return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Meta_Grind_foldProjs___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_foldProjs___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Meta_Grind_foldProjs___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Meta_Grind_foldProjs___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -5784,6 +6316,14 @@ lean_dec(x_2); return x_5; } } +LEAN_EXPORT lean_object* l_Lean_Meta_Grind_normalize___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = lean_grind_normalize(x_1, x_2, x_3, x_4, x_5, x_6); +return x_7; +} +} lean_object* initialize_Lean_Meta_AbstractNestedProofs(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Transform(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Util(uint8_t builtin, lean_object*); @@ -5897,6 +6437,30 @@ l_Lean_Meta_Grind_eraseIrrelevantMData___closed__1 = _init_l_Lean_Meta_Grind_era lean_mark_persistent(l_Lean_Meta_Grind_eraseIrrelevantMData___closed__1); l_Lean_Meta_Grind_eraseIrrelevantMData___closed__2 = _init_l_Lean_Meta_Grind_eraseIrrelevantMData___closed__2(); lean_mark_persistent(l_Lean_Meta_Grind_eraseIrrelevantMData___closed__2); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__1); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__2); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__3); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__4); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__5); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__6); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__7); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__8); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__9); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__10); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__11); +l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12 = _init_l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12(); +lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___lambda__2___closed__12); l_Lean_Meta_Grind_foldProjs___closed__1 = _init_l_Lean_Meta_Grind_foldProjs___closed__1(); lean_mark_persistent(l_Lean_Meta_Grind_foldProjs___closed__1); l_Lean_Meta_Grind_foldProjs___closed__2 = _init_l_Lean_Meta_Grind_foldProjs___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Revert.c b/stage0/stdlib/Lean/Meta/Tactic/Revert.c index c69d87de2796..e569b7d33908 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Revert.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Revert.c @@ -69,7 +69,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_revert___lambda__1(lean_object*, size_t, uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_revert___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MVarId_revertAfter___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -830,7 +830,8 @@ if (x_95 == 0) lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; x_96 = lean_ctor_get(x_94, 0); x_97 = lean_ctor_get(x_94, 1); -x_98 = lean_environment_main_module(x_85); +x_98 = l_Lean_Environment_mainModule(x_85); +lean_dec(x_85); lean_ctor_set(x_94, 1, x_81); lean_ctor_set(x_94, 0, x_98); x_99 = lean_ctor_get(x_96, 1); @@ -1230,7 +1231,8 @@ x_214 = lean_ctor_get(x_94, 1); lean_inc(x_214); lean_inc(x_213); lean_dec(x_94); -x_215 = lean_environment_main_module(x_85); +x_215 = l_Lean_Environment_mainModule(x_85); +lean_dec(x_85); x_216 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_216, 0, x_215); lean_ctor_set(x_216, 1, x_81); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c index 44567a7a9b1a..5046e9ed66e4 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c @@ -209,7 +209,7 @@ static lean_object* l_Lean_Meta_Simp_simpForall___lambda__5___closed__9; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_Simp_simpProj___spec__8(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Simp_simpLet___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpNonDepLetFun_go___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpArrow___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*); @@ -1891,7 +1891,8 @@ x_15 = lean_ctor_get(x_12, 1); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -x_17 = lean_environment_find(x_16, x_11); +x_17 = l_Lean_Environment_find_x3f(x_16, x_11); +lean_dec(x_11); if (lean_obj_tag(x_17) == 0) { lean_object* x_18; @@ -2769,7 +2770,8 @@ lean_dec(x_12); x_245 = lean_ctor_get(x_243, 0); lean_inc(x_245); lean_dec(x_243); -x_246 = lean_environment_find(x_245, x_11); +x_246 = l_Lean_Environment_find_x3f(x_245, x_11); +lean_dec(x_11); if (lean_obj_tag(x_246) == 0) { lean_object* x_247; lean_object* x_248; @@ -4036,7 +4038,8 @@ x_23 = lean_ctor_get(x_20, 1); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_find(x_24, x_2); +x_25 = l_Lean_Environment_find_x3f(x_24, x_2); +lean_dec(x_2); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; @@ -4115,7 +4118,8 @@ lean_dec(x_20); x_40 = lean_ctor_get(x_38, 0); lean_inc(x_40); lean_dec(x_38); -x_41 = lean_environment_find(x_40, x_2); +x_41 = l_Lean_Environment_find_x3f(x_40, x_2); +lean_dec(x_2); if (lean_obj_tag(x_41) == 0) { lean_object* x_42; lean_object* x_43; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c index 44d575d7cc6d..9bd72a9b8e3d 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c @@ -139,7 +139,7 @@ LEAN_EXPORT lean_object* l_Array_insertionSort_traverse___at_Lean_Meta_Simp_rewr LEAN_EXPORT lean_object* l_Lean_Meta_Simp_sevalGround___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ppOrigin___at_Lean_Meta_Simp_discharge_x3f_x27___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_ppSimpTheorem___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__1___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Simp_drewritePre___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*); @@ -971,8 +971,7 @@ x_13 = lean_ctor_get(x_10, 1); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -lean_inc(x_1); -x_15 = lean_environment_find(x_14, x_1); +x_15 = l_Lean_Environment_find_x3f(x_14, x_1); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -1012,8 +1011,7 @@ lean_dec(x_10); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_1); -x_27 = lean_environment_find(x_26, x_1); +x_27 = l_Lean_Environment_find_x3f(x_26, x_1); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c index 2485d032eef8..db12ab2d2fbf 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c @@ -49,7 +49,7 @@ lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at___private_Lean_Meta_Tactic_Simp_SimpAll_0__Lean_Meta_SimpAll_loop___spec__7___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_simpAll___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpAll___hyg_2248_(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Meta_Tactic_Simp_SimpAll_0__Lean_Meta_SimpAll_loop___spec__2___rarg(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_simpAll___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1370,8 +1370,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -1411,8 +1410,7 @@ lean_dec(x_8); x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); lean_dec(x_22); -lean_inc(x_1); -x_25 = lean_environment_find(x_24, x_1); +x_25 = l_Lean_Environment_find_x3f(x_24, x_1); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c index fad84975e4f0..aec3af2a9802 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c @@ -115,7 +115,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocArrayCore___boxed(lean_object*, static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_187____closed__9; static lean_object* l_Lean_Meta_Simp_simprocCore___closed__16; static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_1151____closed__2; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_instInhabitedSimprocDecl___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_187____spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5749____closed__4; @@ -4639,9 +4639,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimprocFromDeclImpl(lean_object* x_ lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); -lean_inc(x_1); lean_inc(x_4); -x_5 = lean_environment_find(x_4, x_1); +x_5 = l_Lean_Environment_find_x3f(x_4, x_1); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Types.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Types.c index 533afd719b6c..ed0feda1b2cb 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Types.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Types.c @@ -140,7 +140,6 @@ static lean_object* l_Lean_Meta_Simp_instInhabitedContext___closed__9; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Simp_tryAutoCongrTheorem_x3f___spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_Result_mkCast___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_removeUnnecessaryCasts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_tryAutoCongrTheorem_x3f___lambda__3___boxed(lean_object**); static lean_object* l_Lean_Meta_Simp_instInhabitedMethods___closed__2; @@ -181,7 +180,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_Meta_Simp_Us LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Simp_congrArgs___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_tryAutoCongrTheorem_x3f___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* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Meta_Simp_mkCongrSimp_x3f___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_tryAutoCongrTheorem_x3f___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instInhabitedMethods___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -370,12 +368,14 @@ uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Simp_recordTriedSimpTheorem___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_Result_mkCast___closed__2; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withFreshCache(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpIndexConfig(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_recordTriedSimpTheorem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Simp_tryAutoCongrTheorem_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Simp_UsedSimps_insert___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Simp_congrArgs___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_EXPORT lean_object* l_Lean_Meta_Simp_MethodsRef_toMethodsImpl(lean_object*); @@ -9684,12 +9684,12 @@ if (x_26 == 0) lean_object* x_27; lean_object* x_28; x_27 = lean_ctor_get(x_22, 2); lean_inc(x_27); -x_28 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_27, x_1); +x_28 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_27, x_1); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; x_29 = lean_unsigned_to_nat(1u); -x_30 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_27, x_1, x_29); +x_30 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_27, x_1, x_29); lean_ctor_set(x_22, 2, x_30); x_31 = lean_st_ref_set(x_4, x_21, x_23); x_32 = !lean_is_exclusive(x_31); @@ -9724,7 +9724,7 @@ lean_dec(x_28); x_39 = lean_unsigned_to_nat(1u); x_40 = lean_nat_add(x_38, x_39); lean_dec(x_38); -x_41 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_27, x_1, x_40); +x_41 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_27, x_1, x_40); lean_ctor_set(x_22, 2, x_41); x_42 = lean_st_ref_set(x_4, x_21, x_23); x_43 = !lean_is_exclusive(x_42); @@ -9764,12 +9764,12 @@ lean_inc(x_50); lean_inc(x_49); lean_dec(x_22); lean_inc(x_51); -x_53 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_51, x_1); +x_53 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_51, x_1); if (lean_obj_tag(x_53) == 0) { lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; x_54 = lean_unsigned_to_nat(1u); -x_55 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_51, x_1, x_54); +x_55 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_51, x_1, x_54); x_56 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_56, 0, x_49); lean_ctor_set(x_56, 1, x_50); @@ -9806,7 +9806,7 @@ lean_dec(x_53); x_63 = lean_unsigned_to_nat(1u); x_64 = lean_nat_add(x_62, x_63); lean_dec(x_62); -x_65 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_51, x_1, x_64); +x_65 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_51, x_1, x_64); x_66 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_66, 0, x_49); lean_ctor_set(x_66, 1, x_50); @@ -9867,12 +9867,12 @@ if (lean_is_exclusive(x_22)) { x_80 = lean_box(0); } lean_inc(x_78); -x_81 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_78, x_1); +x_81 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__1(x_78, x_1); if (lean_obj_tag(x_81) == 0) { lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; x_82 = lean_unsigned_to_nat(1u); -x_83 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_78, x_1, x_82); +x_83 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_78, x_1, x_82); if (lean_is_scalar(x_80)) { x_84 = lean_alloc_ctor(0, 4, 0); } else { @@ -9918,7 +9918,7 @@ lean_dec(x_81); x_92 = lean_unsigned_to_nat(1u); x_93 = lean_nat_add(x_91, x_92); lean_dec(x_91); -x_94 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_78, x_1, x_93); +x_94 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Environment_Diagnostics_recordUnfold___spec__4(x_78, x_1, x_93); if (lean_is_scalar(x_80)) { x_95 = lean_alloc_ctor(0, 4, 0); } else { diff --git a/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c b/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c index 16a1685d25d1..71f81ac13259 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c +++ b/stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c @@ -181,7 +181,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_MVarId_applyRule uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Meta_Iterator_ofList___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__7; static lean_object* l_Lean_Meta_SolveByElim_initFn____x40_Lean_Meta_Tactic_SolveByElim___hyg_5____closed__12; static lean_object* l_Lean_Meta_SolveByElim_solveByElim___closed__1; @@ -7548,7 +7548,8 @@ x_19 = lean_ctor_get(x_16, 1); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); +x_21 = l_Lean_Environment_mainModule(x_20); +lean_dec(x_20); x_22 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__3; lean_inc(x_15); x_23 = l_Lean_addMacroScope(x_21, x_22, x_15); @@ -7571,7 +7572,8 @@ x_31 = lean_ctor_get(x_28, 1); x_32 = lean_ctor_get(x_30, 0); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_environment_main_module(x_32); +x_33 = l_Lean_Environment_mainModule(x_32); +lean_dec(x_32); x_34 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__8; lean_inc(x_15); x_35 = l_Lean_addMacroScope(x_33, x_34, x_15); @@ -7593,7 +7595,8 @@ x_42 = lean_ctor_get(x_39, 1); x_43 = lean_ctor_get(x_41, 0); lean_inc(x_43); lean_dec(x_41); -x_44 = lean_environment_main_module(x_43); +x_44 = l_Lean_Environment_mainModule(x_43); +lean_dec(x_43); x_45 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__13; lean_inc(x_15); x_46 = l_Lean_addMacroScope(x_44, x_45, x_15); @@ -7615,7 +7618,8 @@ x_53 = lean_ctor_get(x_50, 1); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__18; x_57 = l_Lean_addMacroScope(x_55, x_56, x_15); x_58 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__17; @@ -7810,7 +7814,8 @@ lean_dec(x_50); x_106 = lean_ctor_get(x_104, 0); lean_inc(x_106); lean_dec(x_104); -x_107 = lean_environment_main_module(x_106); +x_107 = l_Lean_Environment_mainModule(x_106); +lean_dec(x_106); x_108 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__18; x_109 = l_Lean_addMacroScope(x_107, x_108, x_15); x_110 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__17; @@ -8010,7 +8015,8 @@ lean_dec(x_39); x_159 = lean_ctor_get(x_157, 0); lean_inc(x_159); lean_dec(x_157); -x_160 = lean_environment_main_module(x_159); +x_160 = l_Lean_Environment_mainModule(x_159); +lean_dec(x_159); x_161 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__13; lean_inc(x_15); x_162 = l_Lean_addMacroScope(x_160, x_161, x_15); @@ -8038,7 +8044,8 @@ if (lean_is_exclusive(x_166)) { x_170 = lean_ctor_get(x_167, 0); lean_inc(x_170); lean_dec(x_167); -x_171 = lean_environment_main_module(x_170); +x_171 = l_Lean_Environment_mainModule(x_170); +lean_dec(x_170); x_172 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__18; x_173 = l_Lean_addMacroScope(x_171, x_172, x_15); x_174 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__17; @@ -8243,7 +8250,8 @@ lean_dec(x_28); x_224 = lean_ctor_get(x_222, 0); lean_inc(x_224); lean_dec(x_222); -x_225 = lean_environment_main_module(x_224); +x_225 = l_Lean_Environment_mainModule(x_224); +lean_dec(x_224); x_226 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__8; lean_inc(x_15); x_227 = l_Lean_addMacroScope(x_225, x_226, x_15); @@ -8271,7 +8279,8 @@ if (lean_is_exclusive(x_231)) { x_235 = lean_ctor_get(x_232, 0); lean_inc(x_235); lean_dec(x_232); -x_236 = lean_environment_main_module(x_235); +x_236 = l_Lean_Environment_mainModule(x_235); +lean_dec(x_235); x_237 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__13; lean_inc(x_15); x_238 = l_Lean_addMacroScope(x_236, x_237, x_15); @@ -8299,7 +8308,8 @@ if (lean_is_exclusive(x_242)) { x_246 = lean_ctor_get(x_243, 0); lean_inc(x_246); lean_dec(x_243); -x_247 = lean_environment_main_module(x_246); +x_247 = l_Lean_Environment_mainModule(x_246); +lean_dec(x_246); x_248 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__18; x_249 = l_Lean_addMacroScope(x_247, x_248, x_15); x_250 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__17; @@ -8509,7 +8519,8 @@ lean_dec(x_16); x_301 = lean_ctor_get(x_299, 0); lean_inc(x_301); lean_dec(x_299); -x_302 = lean_environment_main_module(x_301); +x_302 = l_Lean_Environment_mainModule(x_301); +lean_dec(x_301); x_303 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__3; lean_inc(x_15); x_304 = l_Lean_addMacroScope(x_302, x_303, x_15); @@ -8538,7 +8549,8 @@ if (lean_is_exclusive(x_309)) { x_313 = lean_ctor_get(x_310, 0); lean_inc(x_313); lean_dec(x_310); -x_314 = lean_environment_main_module(x_313); +x_314 = l_Lean_Environment_mainModule(x_313); +lean_dec(x_313); x_315 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__8; lean_inc(x_15); x_316 = l_Lean_addMacroScope(x_314, x_315, x_15); @@ -8566,7 +8578,8 @@ if (lean_is_exclusive(x_320)) { x_324 = lean_ctor_get(x_321, 0); lean_inc(x_324); lean_dec(x_321); -x_325 = lean_environment_main_module(x_324); +x_325 = l_Lean_Environment_mainModule(x_324); +lean_dec(x_324); x_326 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__13; lean_inc(x_15); x_327 = l_Lean_addMacroScope(x_325, x_326, x_15); @@ -8594,7 +8607,8 @@ if (lean_is_exclusive(x_331)) { x_335 = lean_ctor_get(x_332, 0); lean_inc(x_335); lean_dec(x_332); -x_336 = lean_environment_main_module(x_335); +x_336 = l_Lean_Environment_mainModule(x_335); +lean_dec(x_335); x_337 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__18; x_338 = l_Lean_addMacroScope(x_336, x_337, x_15); x_339 = l_Lean_Meta_SolveByElim_mkAssumptionSet___lambda__3___closed__17; diff --git a/stage0/stdlib/Lean/Meta/Tactic/TryThis.c b/stage0/stdlib/Lean/Meta/Tactic/TryThis.c index f512a2673ff1..fdc7988d1a41 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/TryThis.c +++ b/stage0/stdlib/Lean/Meta/Tactic/TryThis.c @@ -54,6 +54,7 @@ lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__10; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_tryThisProvider___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__7; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___spec__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -141,13 +142,11 @@ static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__3; static lean_object* l_Lean_Meta_Tactic_TryThis_instInhabitedSuggestionText___closed__1; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_value___closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_addExactSuggestions___spec__1(uint8_t, size_t, size_t, 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_Meta_Tactic_TryThis_getInputWidth___closed__1; static lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addExactSuggestionCore___lambda__3___closed__1; static lean_object* l_Lean_Meta_Tactic_TryThis_delabToRefinableSuggestion___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_Tactic_TryThis_tryThisProvider__1(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___closed__9; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__45; static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addExactSuggestionCore___spec__1___closed__1; @@ -159,7 +158,6 @@ lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__2; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__34; -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_delabToRefinableSuggestion___closed__2; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__12; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__16; @@ -204,6 +202,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_Suggestion_toJsonAndInfoM___ LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_instInhabitedSuggestionText; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_value___closed__9; lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_instInhabitedSuggestion; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_getIndentAndColumn___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addExactSuggestionCore___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -346,7 +345,7 @@ static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__21; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___spec__1___closed__5; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__66; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalConstNoOverloadWithInfo___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_float_to_string(double); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -414,6 +413,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_instCoeHeadTSyntaxConsSyntax lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); double round(double); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__17; static lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__5; @@ -421,7 +421,6 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___closed__13; static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_608____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_addSuggestion___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___closed__5; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_Suggestion_toJsonAndInfoM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -430,6 +429,7 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_Suggestion_toJsonAndInfoM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_Meta_Tactic_TryThis_SuggestionStyle_value___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___spec__2(size_t, size_t, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l___regBuiltin_Lean_Meta_Tactic_TryThis_tryThisProvider__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionText_prettyExtra(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___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*); @@ -1605,7 +1605,7 @@ lean_dec(x_14); x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_Kernel_isDiagnosticsEnabled(x_17); +x_18 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_17); lean_dec(x_17); if (x_18 == 0) { @@ -1655,7 +1655,7 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean x_24 = lean_ctor_get(x_21, 0); x_25 = lean_ctor_get(x_21, 4); lean_dec(x_25); -x_26 = l_Lean_Kernel_enableDiag(x_24, x_13); +x_26 = l_Lean_Kernel_Environment_enableDiag(x_24, x_13); x_27 = l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__5; lean_ctor_set(x_21, 4, x_27); lean_ctor_set(x_21, 0, x_26); @@ -1685,7 +1685,7 @@ lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_dec(x_21); -x_39 = l_Lean_Kernel_enableDiag(x_32, x_13); +x_39 = l_Lean_Kernel_Environment_enableDiag(x_32, x_13); x_40 = l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__5; x_41 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_41, 0, x_39); @@ -3869,13 +3869,13 @@ x_42 = lean_ctor_get(x_40, 1); x_43 = lean_ctor_get(x_40, 0); lean_dec(x_43); x_44 = lean_array_size(x_30); -x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_44, x_24, x_30); +x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_44, x_24, x_30); x_46 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_46, 0, x_45); x_47 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__1; lean_ctor_set(x_40, 1, x_46); lean_ctor_set(x_40, 0, x_47); -x_48 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_32); +x_48 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_32); x_49 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__2; x_50 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_50, 0, x_49); @@ -3959,14 +3959,14 @@ x_79 = lean_ctor_get(x_40, 1); lean_inc(x_79); lean_dec(x_40); x_80 = lean_array_size(x_30); -x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_80, x_24, x_30); +x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_80, x_24, x_30); x_82 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_82, 0, x_81); x_83 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__1; x_84 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_84, 0, x_83); lean_ctor_set(x_84, 1, x_82); -x_85 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_32); +x_85 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_32); x_86 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__2; x_87 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_87, 0, x_86); @@ -4155,7 +4155,7 @@ if (lean_is_exclusive(x_142)) { x_144 = lean_box(0); } x_145 = lean_array_size(x_131); -x_146 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_145, x_125, x_131); +x_146 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_145, x_125, x_131); x_147 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_147, 0, x_146); x_148 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__1; @@ -4166,7 +4166,7 @@ if (lean_is_scalar(x_144)) { } lean_ctor_set(x_149, 0, x_148); lean_ctor_set(x_149, 1, x_147); -x_150 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_133); +x_150 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_133); x_151 = l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__2; x_152 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_152, 0, x_151); @@ -5978,7 +5978,7 @@ lean_dec(x_14); x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_Kernel_isDiagnosticsEnabled(x_17); +x_18 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_17); lean_dec(x_17); if (x_18 == 0) { @@ -6028,7 +6028,7 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean x_24 = lean_ctor_get(x_21, 0); x_25 = lean_ctor_get(x_21, 4); lean_dec(x_25); -x_26 = l_Lean_Kernel_enableDiag(x_24, x_13); +x_26 = l_Lean_Kernel_Environment_enableDiag(x_24, x_13); x_27 = l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__5; lean_ctor_set(x_21, 4, x_27); lean_ctor_set(x_21, 0, x_26); @@ -6058,7 +6058,7 @@ lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_dec(x_21); -x_39 = l_Lean_Kernel_enableDiag(x_32, x_13); +x_39 = l_Lean_Kernel_Environment_enableDiag(x_32, x_13); x_40 = l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__5; x_41 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_41, 0, x_39); @@ -7597,7 +7597,8 @@ x_86 = lean_ctor_get(x_83, 1); x_87 = lean_ctor_get(x_85, 0); lean_inc(x_87); lean_dec(x_85); -x_88 = lean_environment_main_module(x_87); +x_88 = l_Lean_Environment_mainModule(x_87); +lean_dec(x_87); x_89 = l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__18; lean_inc(x_81); lean_ctor_set_tag(x_83, 2); @@ -7656,7 +7657,8 @@ lean_dec(x_83); x_113 = lean_ctor_get(x_111, 0); lean_inc(x_113); lean_dec(x_111); -x_114 = lean_environment_main_module(x_113); +x_114 = l_Lean_Environment_mainModule(x_113); +lean_dec(x_113); x_115 = l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__18; lean_inc(x_81); x_116 = lean_alloc_ctor(2, 2, 0); @@ -8012,7 +8014,8 @@ x_262 = lean_ctor_get(x_259, 1); x_263 = lean_ctor_get(x_261, 0); lean_inc(x_263); lean_dec(x_261); -x_264 = lean_environment_main_module(x_263); +x_264 = l_Lean_Environment_mainModule(x_263); +lean_dec(x_263); x_265 = l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__18; lean_inc(x_257); lean_ctor_set_tag(x_259, 2); @@ -8080,7 +8083,8 @@ lean_dec(x_259); x_294 = lean_ctor_get(x_292, 0); lean_inc(x_294); lean_dec(x_292); -x_295 = lean_environment_main_module(x_294); +x_295 = l_Lean_Environment_mainModule(x_294); +lean_dec(x_294); x_296 = l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__18; lean_inc(x_257); x_297 = lean_alloc_ctor(2, 2, 0); diff --git a/stage0/stdlib/Lean/Meta/Transform.c b/stage0/stdlib/Lean/Meta/Transform.c index 0fd5fbc45732..4c0d47353eb1 100644 --- a/stage0/stdlib/Lean/Meta/Transform.c +++ b/stage0/stdlib/Lean/Meta/Transform.c @@ -107,7 +107,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___rarg___boxed( LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_zetaReduce___spec__16(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTransformStep___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_eraseInaccessibleAnnotations___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_transform_visit___spec__3___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, size_t, lean_object*, lean_object*); @@ -11294,7 +11294,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_unfoldDeclsFrom___lambda__1(lean_object* x_ _start: { lean_object* x_9; -x_9 = lean_environment_find(x_1, x_2); +x_9 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_9) == 0) { lean_object* x_10; lean_object* x_11; @@ -11483,6 +11483,7 @@ if (x_9 == 0) lean_object* x_10; lean_object* x_11; x_10 = lean_box(0); x_11 = l_Lean_Meta_unfoldDeclsFrom___lambda__1(x_1, x_7, x_3, x_8, x_10, x_4, x_5, x_6); +lean_dec(x_7); return x_11; } else @@ -11622,6 +11623,7 @@ x_9 = l_Lean_Meta_unfoldDeclsFrom___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_2); return x_9; } } diff --git a/stage0/stdlib/Lean/Meta/WHNF.c b/stage0/stdlib/Lean/Meta/WHNF.c index 872b3a3a1397..3588f13063ed 100644 --- a/stage0/stdlib/Lean/Meta/WHNF.c +++ b/stage0/stdlib/Lean/Meta/WHNF.c @@ -152,7 +152,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___lambda__2___boxed(lean_o static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__28; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_whnfImp___spec__1___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reducePow___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_Meta_reducePow___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_matchPatternAttr; @@ -1109,7 +1109,8 @@ x_11 = lean_ctor_get(x_9, 0); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); lean_dec(x_11); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); +lean_dec(x_1); lean_ctor_set(x_9, 0, x_13); return x_9; } @@ -1124,7 +1125,8 @@ lean_dec(x_9); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -x_17 = lean_environment_find(x_16, x_1); +x_17 = l_Lean_Environment_find_x3f(x_16, x_1); +lean_dec(x_1); x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_15); @@ -1168,7 +1170,7 @@ x_11 = lean_ctor_get(x_9, 0); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); lean_dec(x_11); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); lean_ctor_set(x_9, 0, x_13); return x_9; } @@ -1183,7 +1185,7 @@ lean_dec(x_9); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -x_17 = lean_environment_find(x_16, x_1); +x_17 = l_Lean_Environment_find_x3f(x_16, x_1); x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_15); @@ -1200,6 +1202,7 @@ x_8 = lean_unbox(x_2); lean_dec(x_2); x_9 = l___private_Lean_Meta_WHNF_0__Lean_Meta_getConstInfoNoEx_x3f(x_1, x_8, x_3, x_4, x_5, x_6, x_7); lean_dec(x_4); +lean_dec(x_1); return x_9; } } @@ -1572,6 +1575,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_getFirstCtor___ lean_object* x_7; x_7 = l___private_Lean_Meta_WHNF_0__Lean_Meta_getFirstCtor(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); +lean_dec(x_1); return x_7; } } @@ -1598,6 +1602,7 @@ x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); x_11 = l___private_Lean_Meta_WHNF_0__Lean_Meta_getFirstCtor(x_9, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_9); if (lean_obj_tag(x_11) == 0) { lean_object* x_12; @@ -4013,6 +4018,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); x_21 = l___private_Lean_Meta_WHNF_0__Lean_Meta_getFirstCtor(x_10, x_5, x_6, x_7, x_8, x_18); +lean_dec(x_10); if (lean_obj_tag(x_21) == 0) { lean_object* x_22; @@ -4207,6 +4213,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); x_61 = l___private_Lean_Meta_WHNF_0__Lean_Meta_getFirstCtor(x_10, x_5, x_6, x_7, x_8, x_58); +lean_dec(x_10); if (lean_obj_tag(x_61) == 0) { lean_object* x_62; @@ -4483,7 +4490,6 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -lean_inc(x_1); x_14 = l_Lean_isStructureLike(x_13, x_1); if (x_14 == 0) { @@ -4491,7 +4497,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_1); lean_ctor_set(x_9, 0, x_2); return x_9; } @@ -4552,7 +4557,6 @@ x_26 = lean_ctor_get(x_24, 0); x_27 = lean_ctor_get(x_24, 1); x_28 = l_Lean_Expr_getAppFn(x_26); x_29 = l_Lean_Expr_isConstOf(x_28, x_1); -lean_dec(x_1); if (x_29 == 0) { lean_dec(x_28); @@ -4583,7 +4587,6 @@ lean_inc(x_32); lean_dec(x_24); x_34 = l_Lean_Expr_getAppFn(x_32); x_35 = l_Lean_Expr_isConstOf(x_34, x_1); -lean_dec(x_1); if (x_35 == 0) { lean_object* x_36; @@ -4615,7 +4618,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -lean_dec(x_1); x_39 = !lean_is_exclusive(x_21); if (x_39 == 0) { @@ -4644,7 +4646,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -lean_dec(x_1); x_43 = !lean_is_exclusive(x_18); if (x_43 == 0) { @@ -4673,7 +4674,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_1); x_47 = !lean_is_exclusive(x_15); if (x_47 == 0) { @@ -4704,7 +4704,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -lean_dec(x_1); x_51 = !lean_is_exclusive(x_15); if (x_51 == 0) { @@ -4737,7 +4736,6 @@ lean_dec(x_9); x_57 = lean_ctor_get(x_55, 0); lean_inc(x_57); lean_dec(x_55); -lean_inc(x_1); x_58 = l_Lean_isStructureLike(x_57, x_1); if (x_58 == 0) { @@ -4746,7 +4744,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_1); x_59 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_59, 0, x_2); lean_ctor_set(x_59, 1, x_56); @@ -4814,7 +4811,6 @@ if (lean_is_exclusive(x_69)) { } x_73 = l_Lean_Expr_getAppFn(x_70); x_74 = l_Lean_Expr_isConstOf(x_73, x_1); -lean_dec(x_1); if (x_74 == 0) { lean_object* x_75; @@ -4850,7 +4846,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -lean_dec(x_1); x_78 = lean_ctor_get(x_66, 0); lean_inc(x_78); x_79 = lean_ctor_get(x_66, 1); @@ -4881,7 +4876,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -lean_dec(x_1); x_82 = lean_ctor_get(x_63, 0); lean_inc(x_82); x_83 = lean_ctor_get(x_63, 1); @@ -4912,7 +4906,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_1); x_86 = lean_ctor_get(x_60, 1); lean_inc(x_86); if (lean_is_exclusive(x_60)) { @@ -4941,7 +4934,6 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -lean_dec(x_1); x_89 = lean_ctor_get(x_60, 0); lean_inc(x_89); x_90 = lean_ctor_get(x_60, 1); @@ -5014,6 +5006,7 @@ lean_inc(x_15); lean_dec(x_8); x_16 = lean_box(0); x_17 = l___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenStructure___lambda__2(x_1, x_2, x_16, x_3, x_4, x_5, x_6, x_15); +lean_dec(x_1); return x_17; } } @@ -5072,6 +5065,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenStruc lean_object* x_9; x_9 = l___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenStructure___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); +lean_dec(x_1); return x_9; } } @@ -6062,6 +6056,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); x_26 = l_Lean_Meta_getUnfoldableConstNoEx_x3f(x_25, x_5, x_6, x_7, x_8, x_23); +lean_dec(x_25); if (lean_obj_tag(x_26) == 0) { lean_object* x_27; @@ -6342,6 +6337,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); x_89 = l_Lean_Meta_getUnfoldableConstNoEx_x3f(x_88, x_5, x_6, x_7, x_8, x_86); +lean_dec(x_88); if (lean_obj_tag(x_89) == 0) { lean_object* x_90; @@ -13279,7 +13275,6 @@ lean_inc(x_36); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -lean_inc(x_36); x_37 = l_Lean_Meta_getUnfoldableConstNoEx_x3f(x_36, x_2, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_37) == 0) { @@ -18383,7 +18378,8 @@ x_13 = lean_ctor_get(x_11, 0); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); lean_dec(x_13); -x_15 = lean_environment_find(x_14, x_10); +x_15 = l_Lean_Environment_find_x3f(x_14, x_10); +lean_dec(x_10); if (lean_obj_tag(x_15) == 0) { lean_object* x_16; @@ -18522,7 +18518,8 @@ lean_dec(x_11); x_47 = lean_ctor_get(x_45, 0); lean_inc(x_47); lean_dec(x_45); -x_48 = lean_environment_find(x_47, x_10); +x_48 = l_Lean_Environment_find_x3f(x_47, x_10); +lean_dec(x_10); if (lean_obj_tag(x_48) == 0) { lean_object* x_49; lean_object* x_50; @@ -24110,7 +24107,6 @@ lean_dec(x_1); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); -lean_inc(x_8); x_10 = l___private_Lean_Meta_WHNF_0__Lean_Meta_getConstInfoNoEx_x3f(x_8, x_2, x_3, x_4, x_5, x_6, x_7); if (lean_obj_tag(x_10) == 0) { @@ -24537,7 +24533,8 @@ x_99 = l_Lean_ConstantInfo_name(x_75); x_100 = l_Lean_Meta_smartUnfoldingSuffix; lean_inc(x_99); x_101 = l_Lean_Name_str___override(x_99, x_100); -x_102 = lean_environment_find(x_98, x_101); +x_102 = l_Lean_Environment_find_x3f(x_98, x_101); +lean_dec(x_101); if (lean_obj_tag(x_102) == 0) { lean_object* x_103; lean_object* x_104; @@ -24941,7 +24938,8 @@ x_198 = l_Lean_ConstantInfo_name(x_172); x_199 = l_Lean_Meta_smartUnfoldingSuffix; lean_inc(x_198); x_200 = l_Lean_Name_str___override(x_198, x_199); -x_201 = lean_environment_find(x_197, x_200); +x_201 = l_Lean_Environment_find_x3f(x_197, x_200); +lean_dec(x_200); if (lean_obj_tag(x_201) == 0) { lean_object* x_202; lean_object* x_203; diff --git a/stage0/stdlib/Lean/Modifiers.c b/stage0/stdlib/Lean/Modifiers.c index 4b567111d687..72ea42d3f087 100644 --- a/stage0/stdlib/Lean/Modifiers.c +++ b/stage0/stdlib/Lean/Modifiers.c @@ -36,11 +36,12 @@ LEAN_EXPORT uint8_t l_Lean_isProtected(lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Modifiers___hyg_3____closed__2; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isPrivateNameFromImportedModule(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isPrivatePrefix_go(lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkPrivateName___boxed(lean_object*, lean_object*); lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Modifiers___hyg_3____closed__1; LEAN_EXPORT uint8_t l_Lean_isPrivateName(lean_object*); @@ -153,7 +154,7 @@ LEAN_EXPORT lean_object* l_Lean_mkPrivateName(lean_object* x_1, lean_object* x_2 _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_3 = lean_environment_main_module(x_1); +x_3 = l_Lean_Environment_mainModule(x_1); x_4 = l_Lean_privateHeader; x_5 = l_Lean_Name_append(x_4, x_3); x_6 = lean_unsigned_to_nat(0u); @@ -162,6 +163,15 @@ x_8 = l_Lean_Name_append(x_7, x_2); return x_8; } } +LEAN_EXPORT lean_object* l_Lean_mkPrivateName___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_mkPrivateName(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} LEAN_EXPORT uint8_t l_Lean_isPrivateName(lean_object* x_1) { _start: { @@ -391,7 +401,6 @@ if (lean_obj_tag(x_3) == 0) { uint8_t x_4; lean_dec(x_2); -lean_dec(x_1); x_4 = 0; return x_4; } @@ -425,6 +434,7 @@ LEAN_EXPORT lean_object* l_Lean_isPrivateNameFromImportedModule___boxed(lean_obj { uint8_t x_3; lean_object* x_4; x_3 = l_Lean_isPrivateNameFromImportedModule(x_1, x_2); +lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } diff --git a/stage0/stdlib/Lean/MonadEnv.c b/stage0/stdlib/Lean/MonadEnv.c index 5352c0796e0c..de040193bfa1 100644 --- a/stage0/stdlib/Lean/MonadEnv.c +++ b/stage0/stdlib/Lean/MonadEnv.c @@ -20,6 +20,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstCtor___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isRec___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConstCheck___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_matchConstCtor___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_allM___at_Lean_isEnumType___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_allM___at_Lean_isEnumType___spec__1___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoRec___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -31,6 +32,7 @@ lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); LEAN_EXPORT lean_object* l_Lean_withoutModifyingEnv___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasConst___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isEnumType___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_matchConstStructure___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasConst___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isInductive___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findModuleOf_x3f___rarg___lambda__1(lean_object*, lean_object*, lean_object*); @@ -40,7 +42,7 @@ static lean_object* l_Lean_getConstInfo___rarg___lambda__1___closed__3; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoDefn___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_evalConst___spec__1(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor(lean_object*); LEAN_EXPORT lean_object* l_Lean_findModuleOf_x3f___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstRec(lean_object*, lean_object*); @@ -65,12 +67,14 @@ static lean_object* l_Lean_getConstInfoInduct___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_evalConst(lean_object*); LEAN_EXPORT lean_object* l_Lean_findModuleOf_x3f___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isInductive___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstRec___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstRec___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setEnv___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConst___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoRec___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_matchConstRec___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstCtor(lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConst___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -98,12 +102,14 @@ LEAN_EXPORT lean_object* l_Lean_withEnv(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoRec___rarg___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_withoutModifyingEnv(lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isRec___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_hasConst(lean_object*); static lean_object* l_Lean_getConstInfoCtor___rarg___lambda__1___closed__1; static lean_object* l_Lean_getConstInfoRec___rarg___lambda__1___closed__1; lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstStructure___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConstCheck___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_matchConstStructureLike___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -116,6 +122,7 @@ static lean_object* l_Lean_getConstInfoDefn___rarg___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_withEnv___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstCtor___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isProp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_matchConstInduct___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findModuleOf_x3f___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstStructureLike___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -153,6 +160,7 @@ LEAN_EXPORT lean_object* l_Lean_matchConstInduct___rarg___lambda__1(lean_object* LEAN_EXPORT lean_object* l_Lean_matchConstInduct___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isRec___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_matchConstStructure___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_matchConst___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_allM___at_Lean_isEnumType___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setEnv___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: @@ -277,7 +285,7 @@ LEAN_EXPORT lean_object* l_Lean_isInductive___rarg___lambda__1(lean_object* x_1, _start: { lean_object* x_4; -x_4 = lean_environment_find(x_3, x_1); +x_4 = l_Lean_Environment_find_x3f(x_3, x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; @@ -340,7 +348,7 @@ lean_inc(x_4); x_5 = lean_ctor_get(x_2, 0); lean_inc(x_5); lean_dec(x_2); -x_6 = lean_alloc_closure((void*)(l_Lean_isInductive___rarg___lambda__1), 3, 2); +x_6 = lean_alloc_closure((void*)(l_Lean_isInductive___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_6, 0, x_3); lean_closure_set(x_6, 1, x_1); x_7 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_5, x_6); @@ -355,11 +363,20 @@ x_2 = lean_alloc_closure((void*)(l_Lean_isInductive___rarg), 3, 0); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_isInductive___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_isInductive___rarg___lambda__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} LEAN_EXPORT uint8_t l_Lean_isRecCore(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_environment_find(x_1, x_2); +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -394,6 +411,7 @@ LEAN_EXPORT lean_object* l_Lean_isRecCore___boxed(lean_object* x_1, lean_object* { uint8_t x_3; lean_object* x_4; x_3 = l_Lean_isRecCore(x_1, x_2); +lean_dec(x_2); x_4 = lean_box(x_3); return x_4; } @@ -423,7 +441,7 @@ lean_inc(x_4); x_5 = lean_ctor_get(x_2, 0); lean_inc(x_5); lean_dec(x_2); -x_6 = lean_alloc_closure((void*)(l_Lean_isRec___rarg___lambda__1), 3, 2); +x_6 = lean_alloc_closure((void*)(l_Lean_isRec___rarg___lambda__1___boxed), 3, 2); lean_closure_set(x_6, 0, x_1); lean_closure_set(x_6, 1, x_3); x_7 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_5, x_6); @@ -438,6 +456,15 @@ x_2 = lean_alloc_closure((void*)(l_Lean_isRec___rarg), 3, 0); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_isRec___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_isRec___rarg___lambda__1(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_withoutModifyingEnv___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -593,7 +620,7 @@ LEAN_EXPORT lean_object* l_Lean_matchConst___rarg___lambda__1(lean_object* x_1, _start: { lean_object* x_6; -x_6 = lean_environment_find(x_5, x_1); +x_6 = l_Lean_Environment_find_x3f(x_5, x_1); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; lean_object* x_8; @@ -632,7 +659,7 @@ lean_dec(x_1); x_9 = lean_ctor_get(x_2, 0); lean_inc(x_9); lean_dec(x_2); -x_10 = lean_alloc_closure((void*)(l_Lean_matchConst___rarg___lambda__1), 5, 4); +x_10 = lean_alloc_closure((void*)(l_Lean_matchConst___rarg___lambda__1___boxed), 5, 4); lean_closure_set(x_10, 0, x_6); lean_closure_set(x_10, 1, x_4); lean_closure_set(x_10, 2, x_5); @@ -661,11 +688,20 @@ x_3 = lean_alloc_closure((void*)(l_Lean_matchConst___rarg), 5, 0); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_matchConst___rarg___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_matchConst___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_matchConstInduct___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = lean_environment_find(x_5, x_1); +x_6 = l_Lean_Environment_find_x3f(x_5, x_1); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; lean_object* x_8; @@ -721,7 +757,7 @@ lean_dec(x_1); x_9 = lean_ctor_get(x_2, 0); lean_inc(x_9); lean_dec(x_2); -x_10 = lean_alloc_closure((void*)(l_Lean_matchConstInduct___rarg___lambda__1), 5, 4); +x_10 = lean_alloc_closure((void*)(l_Lean_matchConstInduct___rarg___lambda__1___boxed), 5, 4); lean_closure_set(x_10, 0, x_6); lean_closure_set(x_10, 1, x_4); lean_closure_set(x_10, 2, x_5); @@ -750,11 +786,20 @@ x_3 = lean_alloc_closure((void*)(l_Lean_matchConstInduct___rarg), 5, 0); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_matchConstInduct___rarg___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_matchConstInduct___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_matchConstCtor___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = lean_environment_find(x_5, x_1); +x_6 = l_Lean_Environment_find_x3f(x_5, x_1); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; lean_object* x_8; @@ -810,7 +855,7 @@ lean_dec(x_1); x_9 = lean_ctor_get(x_2, 0); lean_inc(x_9); lean_dec(x_2); -x_10 = lean_alloc_closure((void*)(l_Lean_matchConstCtor___rarg___lambda__1), 5, 4); +x_10 = lean_alloc_closure((void*)(l_Lean_matchConstCtor___rarg___lambda__1___boxed), 5, 4); lean_closure_set(x_10, 0, x_6); lean_closure_set(x_10, 1, x_4); lean_closure_set(x_10, 2, x_5); @@ -839,11 +884,20 @@ x_3 = lean_alloc_closure((void*)(l_Lean_matchConstCtor___rarg), 5, 0); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_matchConstCtor___rarg___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_matchConstCtor___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_matchConstRec___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = lean_environment_find(x_5, x_1); +x_6 = l_Lean_Environment_find_x3f(x_5, x_1); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; lean_object* x_8; @@ -899,7 +953,7 @@ lean_dec(x_1); x_9 = lean_ctor_get(x_2, 0); lean_inc(x_9); lean_dec(x_2); -x_10 = lean_alloc_closure((void*)(l_Lean_matchConstRec___rarg___lambda__1), 5, 4); +x_10 = lean_alloc_closure((void*)(l_Lean_matchConstRec___rarg___lambda__1___boxed), 5, 4); lean_closure_set(x_10, 0, x_6); lean_closure_set(x_10, 1, x_4); lean_closure_set(x_10, 2, x_5); @@ -928,6 +982,15 @@ x_3 = lean_alloc_closure((void*)(l_Lean_matchConstRec___rarg), 5, 0); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_matchConstRec___rarg___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_matchConstRec___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_hasConst___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -1083,8 +1146,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfo___rarg___lambda__1(lean_object* x_1 _start: { lean_object* x_5; -lean_inc(x_1); -x_5 = lean_environment_find(x_4, x_1); +x_5 = l_Lean_Environment_find_x3f(x_4, x_1); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; @@ -1602,7 +1664,7 @@ LEAN_EXPORT lean_object* l_Lean_matchConstStructure___rarg___lambda__2(lean_obje _start: { lean_object* x_10; -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { lean_object* x_11; lean_object* x_12; @@ -1715,7 +1777,7 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_2, 0); lean_inc(x_10); lean_inc(x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_matchConstStructure___rarg___lambda__2), 9, 8); +x_11 = lean_alloc_closure((void*)(l_Lean_matchConstStructure___rarg___lambda__2___boxed), 9, 8); lean_closure_set(x_11, 0, x_7); lean_closure_set(x_11, 1, x_5); lean_closure_set(x_11, 2, x_1); @@ -1749,11 +1811,20 @@ x_3 = lean_alloc_closure((void*)(l_Lean_matchConstStructure___rarg), 6, 0); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_matchConstStructure___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_matchConstStructure___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} LEAN_EXPORT lean_object* l_Lean_matchConstStructureLike___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { lean_object* x_11; lean_object* x_12; @@ -1908,7 +1979,7 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_2, 0); lean_inc(x_10); lean_inc(x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_matchConstStructureLike___rarg___lambda__1), 9, 8); +x_11 = lean_alloc_closure((void*)(l_Lean_matchConstStructureLike___rarg___lambda__1___boxed), 9, 8); lean_closure_set(x_11, 0, x_7); lean_closure_set(x_11, 1, x_5); lean_closure_set(x_11, 2, x_1); @@ -1942,6 +2013,15 @@ x_3 = lean_alloc_closure((void*)(l_Lean_matchConstStructureLike___rarg), 6, 0); return x_3; } } +LEAN_EXPORT lean_object* l_Lean_matchConstStructureLike___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_matchConstStructureLike___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_evalConst___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { diff --git a/stage0/stdlib/Lean/Parser/Extension.c b/stage0/stdlib/Lean/Parser/Extension.c index 231af616156f..e1aebfdb9341 100644 --- a/stage0/stdlib/Lean/Parser/Extension.c +++ b/stage0/stdlib/Lean/Parser/Extension.c @@ -71,6 +71,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Parser_Ex LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3784____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStackFn___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3858____closed__5; +LEAN_EXPORT lean_object* l_Lean_Parser_resolveParserNameCore_isParser___boxed(lean_object*, lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Parser_getCategory___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getCategory(lean_object*, lean_object*); @@ -140,7 +141,7 @@ lean_object* l_Lean_Parser_ParserState_mkUnexpectedError(lean_object*, lean_obje LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserExtension_mkInitial(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withOpenFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getConstAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5316____closed__16; lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); @@ -196,6 +197,7 @@ LEAN_EXPORT lean_object* l_List_filterMapTR_go___at_Lean_Parser_resolveParserNam LEAN_EXPORT lean_object* l_Lean_Parser_getAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_swap(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__2(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_List_filterMapTR_go___at_Lean_Parser_resolveParserNameCore___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__1; static lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__2___closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); @@ -6203,9 +6205,8 @@ x_5 = lean_ctor_get(x_3, 0); lean_inc(x_5); x_6 = lean_ctor_get(x_3, 1); lean_inc(x_6); -lean_inc(x_1); lean_inc(x_5); -x_7 = lean_environment_find(x_5, x_1); +x_7 = l_Lean_Environment_find_x3f(x_5, x_1); if (lean_obj_tag(x_7) == 0) { 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; @@ -15770,7 +15771,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_resolveParserNameCore_isParser(lean_objec _start: { lean_object* x_3; -x_3 = lean_environment_find(x_1, x_2); +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_3) == 0) { lean_object* x_4; @@ -15977,6 +15978,15 @@ return x_42; } } } +LEAN_EXPORT lean_object* l_Lean_Parser_resolveParserNameCore_isParser___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Parser_resolveParserNameCore_isParser(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Parser_resolveParserNameCore___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { @@ -16030,7 +16040,6 @@ lean_object* x_17; lean_object* x_18; lean_object* x_19; x_17 = lean_ctor_get(x_9, 0); x_18 = lean_ctor_get(x_9, 1); lean_dec(x_18); -lean_inc(x_17); lean_inc(x_1); x_19 = l_Lean_Parser_resolveParserNameCore_isParser(x_1, x_17); if (lean_obj_tag(x_19) == 0) @@ -16103,7 +16112,6 @@ lean_object* x_33; lean_object* x_34; x_33 = lean_ctor_get(x_9, 0); lean_inc(x_33); lean_dec(x_9); -lean_inc(x_33); lean_inc(x_1); x_34 = l_Lean_Parser_resolveParserNameCore_isParser(x_1, x_33); if (lean_obj_tag(x_34) == 0) @@ -16172,7 +16180,6 @@ if (lean_is_exclusive(x_9)) { lean_dec_ref(x_9); x_46 = lean_box(0); } -lean_inc(x_45); lean_inc(x_1); x_47 = l_Lean_Parser_resolveParserNameCore_isParser(x_1, x_45); if (lean_obj_tag(x_47) == 0) @@ -16267,24 +16274,16 @@ else { lean_object* x_5; lean_object* x_6; x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_2, 1); -lean_inc(x_7); -lean_dec(x_2); x_8 = lean_ctor_get(x_5, 0); -lean_inc(x_8); -lean_dec(x_5); -lean_inc(x_8); lean_inc(x_1); x_9 = l_Lean_Parser_resolveParserNameCore_isParser(x_1, x_8); if (lean_obj_tag(x_9) == 0) { -lean_dec(x_8); x_2 = x_7; goto _start; } @@ -16294,6 +16293,7 @@ lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); +lean_inc(x_8); x_12 = lean_alloc_ctor(1, 1, 1); lean_ctor_set(x_12, 0, x_8); x_13 = lean_unbox(x_11); @@ -16308,11 +16308,7 @@ goto _start; else { lean_object* x_16; -lean_dec(x_6); -lean_dec(x_5); x_16 = lean_ctor_get(x_2, 1); -lean_inc(x_16); -lean_dec(x_2); x_2 = x_16; goto _start; } @@ -16408,6 +16404,7 @@ x_8 = lean_box(0); x_9 = l_Lean_Parser_mkParserState___closed__1; lean_inc(x_1); x_10 = l_List_filterMapTR_go___at_Lean_Parser_resolveParserNameCore___spec__2(x_1, x_7, x_9); +lean_dec(x_7); x_11 = l_List_isEmpty___rarg(x_10); if (x_11 == 0) { @@ -16534,6 +16531,15 @@ lean_dec(x_2); return x_9; } } +LEAN_EXPORT lean_object* l_List_filterMapTR_go___at_Lean_Parser_resolveParserNameCore___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_List_filterMapTR_go___at_Lean_Parser_resolveParserNameCore___spec__2(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_resolveParserNameCore___lambda__1___boxed(lean_object* x_1) { _start: { diff --git a/stage0/stdlib/Lean/Parser/Tactic/Doc.c b/stage0/stdlib/Lean/Parser/Tactic/Doc.c index 578c3a87544d..1d076693c4ad 100644 --- a/stage0/stdlib/Lean/Parser/Tactic/Doc.c +++ b/stage0/stdlib/Lean/Parser/Tactic/Doc.c @@ -92,7 +92,7 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Parser_Tactic_Doc_getTactic LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Parser_Tactic_Doc_aliases___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Doc_isTactic___closed__3; static lean_object* l_Lean_Parser_Tactic_Doc_initFn____x40_Lean_Parser_Tactic_Doc___hyg_2831____lambda__3___closed__3; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Parser_Tactic_Doc_allTags___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Doc_initFn____x40_Lean_Parser_Tactic_Doc___hyg_822____closed__4; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Parser_Tactic_Doc_initFn____x40_Lean_Parser_Tactic_Doc___hyg_2831____spec__9___closed__14; @@ -3997,8 +3997,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); lean_dec(x_10); -lean_inc(x_1); -x_13 = lean_environment_find(x_12, x_1); +x_13 = l_Lean_Environment_find_x3f(x_12, x_1); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; lean_object* x_15; @@ -4144,8 +4143,7 @@ lean_dec(x_8); x_51 = lean_ctor_get(x_49, 0); lean_inc(x_51); lean_dec(x_49); -lean_inc(x_1); -x_52 = lean_environment_find(x_51, x_1); +x_52 = l_Lean_Environment_find_x3f(x_51, x_1); if (lean_obj_tag(x_52) == 0) { lean_object* x_53; lean_object* x_54; @@ -8793,8 +8791,7 @@ x_22 = lean_ctor_get(x_19, 1); x_23 = lean_ctor_get(x_21, 0); lean_inc(x_23); lean_dec(x_21); -lean_inc(x_3); -x_24 = lean_environment_find(x_23, x_3); +x_24 = l_Lean_Environment_find_x3f(x_23, x_3); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; @@ -8946,8 +8943,7 @@ lean_dec(x_19); x_62 = lean_ctor_get(x_60, 0); lean_inc(x_62); lean_dec(x_60); -lean_inc(x_3); -x_63 = lean_environment_find(x_62, x_3); +x_63 = l_Lean_Environment_find_x3f(x_62, x_3); if (lean_obj_tag(x_63) == 0) { lean_object* x_64; lean_object* x_65; diff --git a/stage0/stdlib/Lean/ParserCompiler.c b/stage0/stdlib/Lean/ParserCompiler.c index 310dfa5a8fe8..8bf162d0da49 100644 --- a/stage0/stdlib/Lean/ParserCompiler.c +++ b/stage0/stdlib/Lean/ParserCompiler.c @@ -304,6 +304,7 @@ LEAN_EXPORT lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_ParserCompiler_registerParserCompiler___rarg___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__31(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__35___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ParserCompiler_parserNodeKind_x3f___closed__4; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -432,7 +433,6 @@ static lean_object* l_Lean_ParserCompiler_parserNodeKind_x3f___closed__6; static lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___closed__16; static lean_object* l_Lean_ParserCompiler_parserNodeKind_x3f___closed__9; lean_object* l_List_enumFrom___rarg(lean_object*, lean_object*); -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_ParserCompiler_compileParserExpr___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l___private_Lean_Meta_Basic_0__Lean_Meta_Config_toKey(lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_ParserCompiler_compileParserExpr___spec__3___at_Lean_ParserCompiler_compileParserExpr___spec__4___rarg___closed__1; @@ -21871,7 +21871,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -21911,7 +21911,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -22906,7 +22906,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -22946,7 +22946,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -23941,7 +23941,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -23981,7 +23981,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -24976,7 +24976,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -25016,7 +25016,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -26011,7 +26011,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -26051,7 +26051,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -27101,7 +27101,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -27141,7 +27141,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -28136,7 +28136,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -28176,7 +28176,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -29171,7 +29171,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -29211,7 +29211,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -30206,7 +30206,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -30246,7 +30246,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); @@ -31241,7 +31241,7 @@ if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; x_40 = lean_ctor_get(x_38, 0); -x_41 = l_Lean_KernelException_toMessageData(x_40, x_26); +x_41 = l_Lean_Kernel_Exception_toMessageData(x_40, x_26); x_42 = l_Lean_MessageData_toString(x_41, x_35); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -31281,7 +31281,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_38, 0); lean_inc(x_51); lean_dec(x_38); -x_52 = l_Lean_KernelException_toMessageData(x_51, x_26); +x_52 = l_Lean_Kernel_Exception_toMessageData(x_51, x_26); x_53 = l_Lean_MessageData_toString(x_52, x_35); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); diff --git a/stage0/stdlib/Lean/PrettyPrinter.c b/stage0/stdlib/Lean/PrettyPrinter.c index b1de8a42be8c..94d7b5ec4996 100644 --- a/stage0/stdlib/Lean/PrettyPrinter.c +++ b/stage0/stdlib/Lean/PrettyPrinter.c @@ -56,7 +56,7 @@ static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_ static lean_object* l_panic___at_Lean_MessageData_ofConst___spec__1___closed__2; static lean_object* l_Lean_PrettyPrinter_ppTactic___closed__2; static lean_object* l_Lean_PrettyPrinter_ppExprLegacy___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_ppConstNameWithInfos___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfosM___lambda__2___boxed(lean_object*); lean_object* l_Lean_PrettyPrinter_format(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -76,11 +76,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter__ static lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_maybePrependExprSizes___closed__8; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___rarg___lambda__1(lean_object*, lean_object*); uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_ofFormatWithInfosM___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_noContext(lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_1176____closed__10; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppUsing(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_MessageData_ofLazyM___spec__1(lean_object*, lean_object*, size_t, size_t); @@ -108,6 +106,7 @@ static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_ static lean_object* l_Lean_PrettyPrinter_registerParserCompilers___closed__10; extern lean_object* l_Lean_PrettyPrinter_parenthesizerAttribute; lean_object* l_Lean_MessageData_ofFormat(lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_220____closed__3; static lean_object* l_panic___at_Lean_MessageData_ofConst___spec__1___closed__4; static lean_object* l_Lean_PrettyPrinter_ppTerm___closed__2; @@ -242,6 +241,7 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_wi static lean_object* l_panic___at_Lean_MessageData_ofConst___spec__1___closed__1; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_maybePrependExprSizes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1; static lean_object* l_Lean_PrettyPrinter_ppConstNameWithInfos___closed__4; @@ -630,7 +630,7 @@ lean_dec(x_72); x_75 = lean_ctor_get(x_73, 0); lean_inc(x_75); lean_dec(x_73); -x_76 = l_Lean_Kernel_isDiagnosticsEnabled(x_75); +x_76 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_75); lean_dec(x_75); if (x_76 == 0) { @@ -815,7 +815,7 @@ lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean x_82 = lean_ctor_get(x_79, 0); x_83 = lean_ctor_get(x_79, 4); lean_dec(x_83); -x_84 = l_Lean_Kernel_enableDiag(x_82, x_48); +x_84 = l_Lean_Kernel_Environment_enableDiag(x_82, x_48); lean_ctor_set(x_79, 4, x_53); lean_ctor_set(x_79, 0, x_84); x_85 = lean_st_ref_set(x_70, x_79, x_80); @@ -929,7 +929,7 @@ lean_inc(x_109); lean_inc(x_108); lean_inc(x_107); lean_dec(x_79); -x_114 = l_Lean_Kernel_enableDiag(x_107, x_48); +x_114 = l_Lean_Kernel_Environment_enableDiag(x_107, x_48); x_115 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_115, 0, x_114); lean_ctor_set(x_115, 1, x_108); @@ -2132,8 +2132,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_1); -x_12 = lean_environment_find(x_11, x_1); +x_12 = l_Lean_Environment_find_x3f(x_11, x_1); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -2235,8 +2234,7 @@ lean_dec(x_7); x_40 = lean_ctor_get(x_38, 0); lean_inc(x_40); lean_dec(x_38); -lean_inc(x_1); -x_41 = lean_environment_find(x_40, x_1); +x_41 = l_Lean_Environment_find_x3f(x_40, x_1); if (lean_obj_tag(x_41) == 0) { lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; @@ -2428,7 +2426,7 @@ lean_dec(x_22); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -x_26 = l_Lean_Kernel_isDiagnosticsEnabled(x_25); +x_26 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_25); lean_dec(x_25); if (x_26 == 0) { @@ -2482,7 +2480,7 @@ lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean x_32 = lean_ctor_get(x_29, 0); x_33 = lean_ctor_get(x_29, 4); lean_dec(x_33); -x_34 = l_Lean_Kernel_enableDiag(x_32, x_21); +x_34 = l_Lean_Kernel_Environment_enableDiag(x_32, x_21); lean_ctor_set(x_29, 4, x_7); lean_ctor_set(x_29, 0, x_34); x_35 = lean_st_ref_set(x_10, x_29, x_30); @@ -2565,7 +2563,7 @@ lean_inc(x_52); lean_inc(x_51); lean_inc(x_50); lean_dec(x_29); -x_57 = l_Lean_Kernel_enableDiag(x_50, x_21); +x_57 = l_Lean_Kernel_Environment_enableDiag(x_50, x_21); x_58 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_58, 0, x_57); lean_ctor_set(x_58, 1, x_51); @@ -2761,7 +2759,7 @@ lean_dec(x_110); x_113 = lean_ctor_get(x_111, 0); lean_inc(x_113); lean_dec(x_111); -x_114 = l_Lean_Kernel_isDiagnosticsEnabled(x_113); +x_114 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_113); lean_dec(x_113); if (x_114 == 0) { @@ -2836,7 +2834,7 @@ if (lean_is_exclusive(x_117)) { lean_dec_ref(x_117); x_126 = lean_box(0); } -x_127 = l_Lean_Kernel_enableDiag(x_119, x_109); +x_127 = l_Lean_Kernel_Environment_enableDiag(x_119, x_109); if (lean_is_scalar(x_126)) { x_128 = lean_alloc_ctor(0, 8, 0); } else { @@ -3144,7 +3142,7 @@ lean_dec(x_79); x_82 = lean_ctor_get(x_80, 0); lean_inc(x_82); lean_dec(x_80); -x_83 = l_Lean_Kernel_isDiagnosticsEnabled(x_82); +x_83 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_82); lean_dec(x_82); if (x_83 == 0) { @@ -3334,7 +3332,7 @@ x_89 = lean_ctor_get(x_86, 0); x_90 = lean_ctor_get(x_86, 4); lean_dec(x_90); x_91 = l_Lean_PrettyPrinter_ppExprLegacy___closed__7; -x_92 = l_Lean_Kernel_enableDiag(x_89, x_91); +x_92 = l_Lean_Kernel_Environment_enableDiag(x_89, x_91); lean_ctor_set(x_86, 4, x_61); lean_ctor_set(x_86, 0, x_92); x_93 = lean_st_ref_set(x_77, x_86, x_87); @@ -3449,7 +3447,7 @@ lean_inc(x_116); lean_inc(x_115); lean_dec(x_86); x_122 = l_Lean_PrettyPrinter_ppExprLegacy___closed__7; -x_123 = l_Lean_Kernel_enableDiag(x_115, x_122); +x_123 = l_Lean_Kernel_Environment_enableDiag(x_115, x_122); x_124 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_124, 0, x_123); lean_ctor_set(x_124, 1, x_116); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c index 098959999d61..cf248c1d6128 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c @@ -132,7 +132,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute___lam static lean_object* l_Lean_PrettyPrinter_Delaborator_delab___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delab___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAtomic(lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadWithReaderOfSubExprDelabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_instAlternativeDelabM___closed__3; @@ -140,7 +139,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadStateOfHoleIt extern lean_object* l_Lean_SubExpr_Pos_maxChildren; static lean_object* l_Lean_PrettyPrinter_Delaborator_getUnusedName___closed__2; static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__3; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_PrettyPrinter_Delaborator_addTermInfo___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_shouldOmitProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -202,6 +200,7 @@ extern lean_object* l_Lean_SubExpr_Pos_typeCoord; static lean_object* l_Lean_PrettyPrinter_Delaborator_getExprKind___closed__15; lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_ReaderT_instMonadLift(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); uint8_t l_Lean_getPPInstantiateMVars(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_attrApp__delab_____closed__7; static lean_object* l_Lean_PrettyPrinter_Delaborator_OmissionReason_toString___closed__3; @@ -434,6 +433,7 @@ lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___s static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_4658____closed__11; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_annotateCurPos___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); uint8_t l_Lean_Expr_isMData(lean_object*); static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_4658____closed__9; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_mkAnnotatedIdent___spec__1___boxed(lean_object*); @@ -12174,7 +12174,7 @@ lean_object* x_19; uint8_t x_20; uint8_t x_21; x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = l_Lean_Kernel_isDiagnosticsEnabled(x_19); +x_20 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_19); lean_dec(x_19); if (x_20 == 0) { @@ -12228,7 +12228,7 @@ lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean x_26 = lean_ctor_get(x_23, 0); x_27 = lean_ctor_get(x_23, 4); lean_dec(x_27); -x_28 = l_Lean_Kernel_enableDiag(x_26, x_14); +x_28 = l_Lean_Kernel_Environment_enableDiag(x_26, x_14); x_29 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__4; lean_ctor_set(x_23, 4, x_29); lean_ctor_set(x_23, 0, x_28); @@ -12258,7 +12258,7 @@ lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_dec(x_23); -x_41 = l_Lean_Kernel_enableDiag(x_34, x_14); +x_41 = l_Lean_Kernel_Environment_enableDiag(x_34, x_14); x_42 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__4; x_43 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_43, 0, x_41); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c index 740777c4bdeb..92bb71ae41ed 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.PrettyPrinter.Delaborator.Builtins -// Imports: Lean.Parser Lean.PrettyPrinter.Delaborator.Attributes Lean.PrettyPrinter.Delaborator.Basic Lean.PrettyPrinter.Delaborator.SubExpr Lean.PrettyPrinter.Delaborator.TopDownAnalyze Lean.Meta.CoeAttr +// Imports: Lean.PrettyPrinter.Delaborator.Attributes Lean.PrettyPrinter.Delaborator.Basic Lean.PrettyPrinter.Delaborator.SubExpr Lean.PrettyPrinter.Delaborator.TopDownAnalyze Lean.Meta.CoeAttr #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -359,7 +359,7 @@ lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPr lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__6___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNamedPattern__1___closed__4; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVarAux___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -514,7 +514,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedApp static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetE__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr__8(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); extern lean_object* l_instInhabitedNat; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getNumArgs(lean_object*); @@ -530,7 +529,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore__ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__7(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabBinders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabForallBinders___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_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -749,6 +747,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lamb static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Kernel_Environment_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__3(lean_object*); @@ -1286,7 +1285,7 @@ LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__8; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__2; @@ -1635,6 +1634,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_withoutParentProjections___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Environment_enableDiag(lean_object*, uint8_t); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__2___closed__1; uint8_t l_Lean_Expr_isMData(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withMDataOptions(lean_object*); @@ -2390,7 +2390,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__2; -x_3 = lean_unsigned_to_nat(40u); +x_3 = lean_unsigned_to_nat(39u); x_4 = lean_unsigned_to_nat(35u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2756,7 +2756,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__3 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__2; -x_3 = lean_unsigned_to_nat(51u); +x_3 = lean_unsigned_to_nat(50u); x_4 = lean_unsigned_to_nat(32u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3427,7 +3427,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__2 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__1; -x_3 = lean_unsigned_to_nat(71u); +x_3 = lean_unsigned_to_nat(70u); x_4 = lean_unsigned_to_nat(30u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3539,7 +3539,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabSort___closed__2 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__1; -x_3 = lean_unsigned_to_nat(76u); +x_3 = lean_unsigned_to_nat(75u); x_4 = lean_unsigned_to_nat(30u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -5850,7 +5850,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConst___closed__ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__1; -x_3 = lean_unsigned_to_nat(93u); +x_3 = lean_unsigned_to_nat(92u); x_4 = lean_unsigned_to_nat(35u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -5997,6 +5997,7 @@ lean_inc(x_38); lean_dec(x_36); lean_inc(x_29); x_39 = l_Lean_mkPrivateName(x_38, x_29); +lean_dec(x_38); x_40 = lean_name_eq(x_11, x_39); lean_dec(x_39); if (x_40 == 0) @@ -17262,7 +17263,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__1; -x_3 = lean_unsigned_to_nat(479u); +x_3 = lean_unsigned_to_nat(478u); x_4 = lean_unsigned_to_nat(58u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -23766,8 +23767,7 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -lean_inc(x_1); -x_14 = lean_environment_find(x_13, x_1); +x_14 = l_Lean_Environment_find_x3f(x_13, x_1); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -23807,8 +23807,7 @@ lean_dec(x_9); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -lean_inc(x_1); -x_26 = lean_environment_find(x_25, x_1); +x_26 = l_Lean_Environment_find_x3f(x_25, x_1); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; @@ -26039,7 +26038,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambd lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__2; -x_3 = lean_unsigned_to_nat(710u); +x_3 = lean_unsigned_to_nat(709u); x_4 = lean_unsigned_to_nat(10u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -35371,7 +35370,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__2 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__1; -x_3 = lean_unsigned_to_nat(960u); +x_3 = lean_unsigned_to_nat(959u); x_4 = lean_unsigned_to_nat(38u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -36247,7 +36246,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLit___closed__2( lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabLit___closed__1; -x_3 = lean_unsigned_to_nat(980u); +x_3 = lean_unsigned_to_nat(979u); x_4 = lean_unsigned_to_nat(29u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -39387,7 +39386,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__2 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__1; -x_3 = lean_unsigned_to_nat(1088u); +x_3 = lean_unsigned_to_nat(1087u); x_4 = lean_unsigned_to_nat(36u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -44529,7 +44528,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__4; -x_3 = lean_unsigned_to_nat(1224u); +x_3 = lean_unsigned_to_nat(1223u); x_4 = lean_unsigned_to_nat(40u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -48605,7 +48604,7 @@ lean_dec(x_14); x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_Kernel_isDiagnosticsEnabled(x_17); +x_18 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_17); lean_dec(x_17); if (x_18 == 0) { @@ -48655,7 +48654,7 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean x_24 = lean_ctor_get(x_21, 0); x_25 = lean_ctor_get(x_21, 4); lean_dec(x_25); -x_26 = l_Lean_Kernel_enableDiag(x_24, x_13); +x_26 = l_Lean_Kernel_Environment_enableDiag(x_24, x_13); x_27 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5; lean_ctor_set(x_21, 4, x_27); lean_ctor_set(x_21, 0, x_26); @@ -48685,7 +48684,7 @@ lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_dec(x_21); -x_39 = l_Lean_Kernel_enableDiag(x_32, x_13); +x_39 = l_Lean_Kernel_Environment_enableDiag(x_32, x_13); x_40 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5; x_41 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_41, 0, x_39); @@ -48739,7 +48738,8 @@ x_7 = lean_ctor_get(x_5, 0); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_main_module(x_8); +x_9 = l_Lean_Environment_mainModule(x_8); +lean_dec(x_8); x_10 = lean_ctor_get(x_2, 10); lean_inc(x_10); lean_dec(x_2); @@ -48758,7 +48758,8 @@ lean_dec(x_5); x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -x_15 = lean_environment_main_module(x_14); +x_15 = l_Lean_Environment_mainModule(x_14); +lean_dec(x_14); x_16 = lean_ctor_get(x_2, 10); lean_inc(x_16); lean_dec(x_2); @@ -53423,7 +53424,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatu lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__1; -x_3 = lean_unsigned_to_nat(1378u); +x_3 = lean_unsigned_to_nat(1377u); x_4 = lean_unsigned_to_nat(42u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -54158,7 +54159,7 @@ lean_dec(x_17); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = l_Lean_Kernel_isDiagnosticsEnabled(x_20); +x_21 = l_Lean_Kernel_Environment_isDiagnosticsEnabled(x_20); lean_dec(x_20); if (x_21 == 0) { @@ -54208,7 +54209,7 @@ lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean x_27 = lean_ctor_get(x_24, 0); x_28 = lean_ctor_get(x_24, 4); lean_dec(x_28); -x_29 = l_Lean_Kernel_enableDiag(x_27, x_16); +x_29 = l_Lean_Kernel_Environment_enableDiag(x_27, x_16); x_30 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5; lean_ctor_set(x_24, 4, x_30); lean_ctor_set(x_24, 0, x_29); @@ -54238,7 +54239,7 @@ lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); lean_dec(x_24); -x_42 = l_Lean_Kernel_enableDiag(x_35, x_16); +x_42 = l_Lean_Kernel_Environment_enableDiag(x_35, x_16); x_43 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5; x_44 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_44, 0, x_42); @@ -54414,7 +54415,6 @@ x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature(x_9, x_2, x_3, x return x_10; } } -lean_object* initialize_Lean_Parser(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter_Delaborator_Attributes(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter_Delaborator_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_PrettyPrinter_Delaborator_SubExpr(uint8_t builtin, lean_object*); @@ -54425,9 +54425,6 @@ LEAN_EXPORT lean_object* initialize_Lean_PrettyPrinter_Delaborator_Builtins(uint lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Lean_Parser(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_PrettyPrinter_Delaborator_Attributes(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c index d81460c21f0e..f1ba168e990c 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c @@ -33,7 +33,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__6___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__2___closed__1; static lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_projInfo___closed__3; @@ -292,7 +292,8 @@ lean_object* x_24; lean_object* x_25; x_24 = lean_ctor_get(x_20, 0); lean_inc(x_24); lean_inc(x_14); -x_25 = lean_environment_find(x_14, x_24); +x_25 = l_Lean_Environment_find_x3f(x_14, x_24); +lean_dec(x_24); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; diff --git a/stage0/stdlib/Lean/PrettyPrinter/Formatter.c b/stage0/stdlib/Lean/PrettyPrinter/Formatter.c index b24040787530..880770f3c3f1 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Formatter.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Formatter.c @@ -116,7 +116,7 @@ static lean_object* l_Lean_PrettyPrinter_Formatter_parseToken___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Formatter_optionalNoAntiquot_formatter___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_formatterForKindUnsafe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_throwBacktrack(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_PrettyPrinter_Formatter_manyNoAntiquot_formatter___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1485,9 +1485,8 @@ goto block_29; else { lean_object* x_31; -lean_inc(x_14); lean_inc(x_10); -x_31 = lean_environment_find(x_10, x_14); +x_31 = l_Lean_Environment_find_x3f(x_10, x_14); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; diff --git a/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c b/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c index db5085170933..c00c7a05415c 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c @@ -129,7 +129,7 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambd LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColGt_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___closed__10; static lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__8; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_PrettyPrinter_Parenthesizer_manyNoAntiquot_parenthesizer___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkKind___closed__7; @@ -1512,9 +1512,8 @@ goto block_29; else { lean_object* x_31; -lean_inc(x_14); lean_inc(x_10); -x_31 = lean_environment_find(x_10, x_14); +x_31 = l_Lean_Environment_find_x3f(x_10, x_14); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; diff --git a/stage0/stdlib/Lean/ProjFns.c b/stage0/stdlib/Lean/ProjFns.c index a8332518c2b3..6b63be702eb7 100644 --- a/stage0/stdlib/Lean/ProjFns.c +++ b/stage0/stdlib/Lean/ProjFns.c @@ -24,7 +24,7 @@ static lean_object* l___private_Lean_ProjFns_0__Lean_reprProjectionFunctionInfo_ LEAN_EXPORT lean_object* l_Lean_getProjectionFnInfo_x3f___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instReprProjectionFunctionInfo; static lean_object* l___private_Lean_ProjFns_0__Lean_reprProjectionFunctionInfo____x40_Lean_ProjFns___hyg_52____closed__16; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l___private_Lean_ProjFns_0__Lean_reprProjectionFunctionInfo____x40_Lean_ProjFns___hyg_52____closed__12; uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_ProjFns___hyg_193_(lean_object*); @@ -734,7 +734,8 @@ lean_dec(x_5); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_environment_find(x_1, x_8); +x_9 = l_Lean_Environment_find_x3f(x_1, x_8); +lean_dec(x_8); if (lean_obj_tag(x_9) == 0) { lean_object* x_10; diff --git a/stage0/stdlib/Lean/ReducibilityAttrs.c b/stage0/stdlib/Lean/ReducibilityAttrs.c index 1db627cdd7cf..2301c1920224 100644 --- a/stage0/stdlib/Lean/ReducibilityAttrs.c +++ b/stage0/stdlib/Lean/ReducibilityAttrs.c @@ -65,7 +65,7 @@ LEAN_EXPORT lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_validate___l lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_addAttr___closed__1; static lean_object* l_Lean_getConstInfo___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__1___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_reprReducibilityStatus____x40_Lean_ReducibilityAttrs___hyg_18____closed__3; static lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_reprReducibilityStatus____x40_Lean_ReducibilityAttrs___hyg_18____closed__16; @@ -3437,8 +3437,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -3478,8 +3477,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; diff --git a/stage0/stdlib/Lean/Replay.c b/stage0/stdlib/Lean/Replay.c index cc76eec0e62f..43e33c6f0380 100644 --- a/stage0/stdlib/Lean/Replay.c +++ b/stage0/stdlib/Lean/Replay.c @@ -13,16 +13,10 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__15; LEAN_EXPORT lean_object* l_Lean_Environment_Replay_replayConstant___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__3; LEAN_EXPORT lean_object* l_Lean_Environment_Replay_replayConstant(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_Replay_replayConstant___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__12; -lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___lambda__1___boxed(lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__7; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Environment_Replay_replayConstant___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); @@ -40,29 +34,20 @@ lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; lean_object* l_Lean_ConstantInfo_name(lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__14; lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__6; lean_object* l_Lean_RBNode_balLeft___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_replay___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_io_get_num_heartbeats(lean_object*); uint8_t l_Lean_ConstantInfo_isUnsafe(lean_object*); lean_object* l_Lean_RBNode_balRight___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_EStateM_instMonad(lean_object*, lean_object*); -extern lean_object* l_Lean_maxRecDepth; -static lean_object* l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Environment_Replay_isTodo___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Environment_Replay_replayConstant___closed__2; -static uint8_t l_Lean_Environment_Replay_throwKernelException___closed__21; -lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); extern lean_object* l_instInhabitedPUnit; -uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldrM___at_Lean_Environment_replay___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_Replay_replayConstant___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__5; +uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_replay(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___rarg(lean_object*, lean_object*); @@ -74,72 +59,53 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Environment_Replay_isTodo static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Environment_Replay_replayConstant___spec__2___closed__1; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); -lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__3; static lean_object* l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__1; -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__13; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Environment_Replay_replayConstant___spec__5(uint64_t, uint64_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_addDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___lambda__1(lean_object*); static lean_object* l_Lean_Environment_Replay_replayConstant___closed__4; lean_object* lean_st_ref_get(lean_object*, lean_object*); -uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Environment_Replay_replayConstant___spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_Replay_replayConstant___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__9; -lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__18; uint8_t lean_name_eq(lean_object*, lean_object*); -lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__1; -uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); -lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1___boxed(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_diagnostics; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__1; static lean_object* l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Environment_Replay_checkPostponedRecursors(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_Replay_replayConstant___spec__7___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_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_replayConstants___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__17; LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Environment_Replay_isTodo___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Environment_Replay_isTodo___spec__2___boxed(lean_object*, lean_object*); +lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(lean_object*, lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Environment_Replay_replayConstant___spec__2___closed__4; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Environment_Replay_replayConstant___spec__9(lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__4; extern lean_object* l_Lean_NameSet_empty; LEAN_EXPORT lean_object* l_Lean_Environment_Replay_checkPostponedConstructors___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Environment_replay___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__16; extern lean_object* l_Lean_instInhabitedConstantInfo; -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__19; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_Replay_replayConstant___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_Exception_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); -LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__2; uint8_t l_Lean_RBNode_isBlack___rarg(lean_object*); +lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_find_x3f___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Environment_Replay_replayConstant___spec__2(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__20; lean_object* l_List_reverse___rarg(lean_object*); -extern lean_object* l_Lean_firstFrontendMacroScope; size_t lean_usize_sub(size_t, size_t); -lean_object* lean_array_mk(lean_object*); +uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__8; uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Environment_Replay_replayConstant___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_replayConstants(lean_object*, lean_object*, lean_object*, lean_object*); @@ -149,7 +115,6 @@ static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Environm lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_Replay_replayConstant___closed__3; -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__11; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Environment_Replay_replayConstant___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -158,11 +123,9 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Environment_Replay_repl LEAN_EXPORT lean_object* l_panic___at_Lean_Environment_Replay_replayConstant___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Environment_replay___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Environment_Replay_replayConstant___spec__2___closed__3; -lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Environment_Replay_replayConstant___spec__3(uint64_t, uint64_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2; uint8_t l_Lean_ConstantInfo_isPartial(lean_object*); -static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__10; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_checkPostponedConstructors(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_isTodo(lean_object*, lean_object*, lean_object*, lean_object*); @@ -577,581 +540,39 @@ lean_dec(x_2); return x_5; } } -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_maxRecDepth; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_5); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_5, 4); -lean_dec(x_9); -x_10 = lean_ctor_get(x_5, 2); -lean_dec(x_10); -x_11 = l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1; -x_12 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_11); -lean_ctor_set(x_5, 4, x_12); -lean_ctor_set(x_5, 2, x_1); -lean_ctor_set_uint8(x_5, sizeof(void*)*12, x_2); -x_13 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_3, x_5, x_6, x_7); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_14 = lean_ctor_get(x_5, 0); -x_15 = lean_ctor_get(x_5, 1); -x_16 = lean_ctor_get(x_5, 3); -x_17 = lean_ctor_get(x_5, 5); -x_18 = lean_ctor_get(x_5, 6); -x_19 = lean_ctor_get(x_5, 7); -x_20 = lean_ctor_get(x_5, 8); -x_21 = lean_ctor_get(x_5, 9); -x_22 = lean_ctor_get(x_5, 10); -x_23 = lean_ctor_get(x_5, 11); -x_24 = lean_ctor_get_uint8(x_5, sizeof(void*)*12 + 1); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_5); -x_25 = l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1; -x_26 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_25); -x_27 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_27, 0, x_14); -lean_ctor_set(x_27, 1, x_15); -lean_ctor_set(x_27, 2, x_1); -lean_ctor_set(x_27, 3, x_16); -lean_ctor_set(x_27, 4, x_26); -lean_ctor_set(x_27, 5, x_17); -lean_ctor_set(x_27, 6, x_18); -lean_ctor_set(x_27, 7, x_19); -lean_ctor_set(x_27, 8, x_20); -lean_ctor_set(x_27, 9, x_21); -lean_ctor_set(x_27, 10, x_22); -lean_ctor_set(x_27, 11, x_23); -lean_ctor_set_uint8(x_27, sizeof(void*)*12, x_2); -lean_ctor_set_uint8(x_27, sizeof(void*)*12 + 1, x_24); -x_28 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_3, x_27, x_6, x_7); -return x_28; -} -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("internal exception #", 20, 20); -return x_1; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Environment_Replay_throwKernelException___closed__3; -x_2 = l_Lean_Environment_Replay_throwKernelException___closed__2; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = l_Lean_Core_getMaxHeartbeats(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_firstFrontendMacroScope; -x_2 = lean_unsigned_to_nat(1u); -x_3 = lean_nat_add(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_uniq", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Environment_Replay_throwKernelException___closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Environment_Replay_throwKernelException___closed__8; -x_2 = lean_unsigned_to_nat(1u); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(32u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Environment_Replay_throwKernelException___closed__10; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__12() { -_start: -{ -size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = 5; -x_2 = l_Lean_Environment_Replay_throwKernelException___closed__11; -x_3 = l_Lean_Environment_Replay_throwKernelException___closed__10; -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_4); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_usize(x_5, 4, x_1); -return x_5; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__13() { -_start: -{ -uint64_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 0; -x_2 = l_Lean_Environment_Replay_throwKernelException___closed__12; -x_3 = lean_alloc_ctor(0, 1, 8); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set_uint64(x_3, sizeof(void*)*1, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__14() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); -return x_1; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Environment_Replay_throwKernelException___closed__14; -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_Environment_Replay_throwKernelException___closed__16() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Environment_Replay_throwKernelException___closed__15; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Environment_Replay_throwKernelException___closed__12; -x_2 = l_Lean_NameSet_empty; -x_3 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_1); -lean_ctor_set(x_3, 2, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__18() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 1; -x_2 = l_Lean_Environment_Replay_throwKernelException___closed__15; -x_3 = l_Lean_Environment_Replay_throwKernelException___closed__12; -x_4 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); -return x_4; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___closed__20() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_diagnostics; -return x_1; -} -} -static uint8_t _init_l_Lean_Environment_Replay_throwKernelException___closed__21() { -_start: -{ -lean_object* x_1; lean_object* x_2; uint8_t x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Environment_Replay_throwKernelException___closed__20; -x_3 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; lean_object* x_11; lean_object* x_12; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; uint8_t x_58; -x_23 = lean_box(0); -x_24 = lean_box(0); -x_25 = lean_st_ref_get(x_3, x_4); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_26, 0); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Environment_Replay_throwKernelException___closed__6; -x_30 = l_Lean_Environment_Replay_throwKernelException___closed__9; -x_31 = l_Lean_Environment_Replay_throwKernelException___closed__13; -x_32 = l_Lean_Environment_Replay_throwKernelException___closed__16; -x_33 = l_Lean_Environment_Replay_throwKernelException___closed__17; -x_34 = l_Lean_Environment_Replay_throwKernelException___closed__18; -x_35 = l_Lean_Environment_Replay_throwKernelException___closed__19; -x_36 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_36, 0, x_28); -lean_ctor_set(x_36, 1, x_29); -lean_ctor_set(x_36, 2, x_30); -lean_ctor_set(x_36, 3, x_31); -lean_ctor_set(x_36, 4, x_32); -lean_ctor_set(x_36, 5, x_33); -lean_ctor_set(x_36, 6, x_34); -lean_ctor_set(x_36, 7, x_35); -x_37 = lean_io_get_num_heartbeats(x_27); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = l_Lean_Environment_Replay_throwKernelException___closed__3; -x_41 = l_Lean_Environment_Replay_throwKernelException___closed__4; -x_42 = lean_unsigned_to_nat(0u); -x_43 = lean_unsigned_to_nat(1000u); -x_44 = lean_box(0); -x_45 = lean_box(0); -x_46 = l_Lean_Environment_Replay_throwKernelException___closed__5; -x_47 = l_Lean_firstFrontendMacroScope; -x_48 = 0; -x_49 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_49, 0, x_40); -lean_ctor_set(x_49, 1, x_41); -lean_ctor_set(x_49, 2, x_23); -lean_ctor_set(x_49, 3, x_42); -lean_ctor_set(x_49, 4, x_43); -lean_ctor_set(x_49, 5, x_44); -lean_ctor_set(x_49, 6, x_45); -lean_ctor_set(x_49, 7, x_23); -lean_ctor_set(x_49, 8, x_38); -lean_ctor_set(x_49, 9, x_46); -lean_ctor_set(x_49, 10, x_47); -lean_ctor_set(x_49, 11, x_24); -lean_ctor_set_uint8(x_49, sizeof(void*)*12, x_48); -lean_ctor_set_uint8(x_49, sizeof(void*)*12 + 1, x_48); -x_50 = lean_st_mk_ref(x_36, x_39); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); -x_53 = lean_st_ref_get(x_51, x_52); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_ctor_get(x_54, 0); -lean_inc(x_56); -lean_dec(x_54); -x_57 = l_Lean_Kernel_isDiagnosticsEnabled(x_56); -lean_dec(x_56); -if (x_57 == 0) -{ -uint8_t x_95; -x_95 = l_Lean_Environment_Replay_throwKernelException___closed__21; -if (x_95 == 0) -{ -uint8_t x_96; -x_96 = 1; -x_58 = x_96; -goto block_94; -} -else -{ -x_58 = x_48; -goto block_94; -} -} -else -{ -uint8_t x_97; -x_97 = l_Lean_Environment_Replay_throwKernelException___closed__21; -if (x_97 == 0) -{ -x_58 = x_48; -goto block_94; -} -else -{ -uint8_t x_98; -x_98 = 1; -x_58 = x_98; -goto block_94; -} -} -block_10: -{ -uint8_t x_6; -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(1, 2, 0); -lean_ctor_set(x_9, 0, x_7); -lean_ctor_set(x_9, 1, x_8); -return x_9; -} -} -block_22: -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_MessageData_toString(x_13, x_12); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_tag(x_14, 1); -lean_ctor_set(x_14, 0, x_17); -x_5 = x_14; -goto block_10; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_14, 0); -x_19 = lean_ctor_get(x_14, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_14); -x_20 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_20, 0, x_18); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -x_5 = x_21; -goto block_10; -} -} -block_94: -{ -if (x_58 == 0) -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; -x_59 = lean_st_ref_take(x_51, x_55); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = !lean_is_exclusive(x_60); -if (x_62 == 0) -{ -lean_object* x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_63 = lean_ctor_get(x_60, 0); -x_64 = lean_ctor_get(x_60, 4); -lean_dec(x_64); -x_65 = l_Lean_Environment_Replay_throwKernelException___closed__21; -x_66 = l_Lean_Kernel_enableDiag(x_63, x_65); -lean_ctor_set(x_60, 4, x_32); -lean_ctor_set(x_60, 0, x_66); -x_67 = lean_st_ref_set(x_51, x_60, x_61); -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -lean_dec(x_67); -x_69 = lean_box(0); -x_70 = l_Lean_Environment_Replay_throwKernelException___lambda__1(x_23, x_65, x_1, x_69, x_49, x_51, x_68); -lean_dec(x_51); -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_11 = x_71; -x_12 = x_72; -goto block_22; -} -else +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_box(0); +x_6 = l_Lean_Kernel_Exception_toMessageData(x_1, x_5); +x_7 = l_Lean_MessageData_toString(x_6, x_4); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; 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_73 = lean_ctor_get(x_60, 0); -x_74 = lean_ctor_get(x_60, 1); -x_75 = lean_ctor_get(x_60, 2); -x_76 = lean_ctor_get(x_60, 3); -x_77 = lean_ctor_get(x_60, 5); -x_78 = lean_ctor_get(x_60, 6); -x_79 = lean_ctor_get(x_60, 7); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_inc(x_76); -lean_inc(x_75); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_60); -x_80 = l_Lean_Environment_Replay_throwKernelException___closed__21; -x_81 = l_Lean_Kernel_enableDiag(x_73, x_80); -x_82 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_74); -lean_ctor_set(x_82, 2, x_75); -lean_ctor_set(x_82, 3, x_76); -lean_ctor_set(x_82, 4, x_32); -lean_ctor_set(x_82, 5, x_77); -lean_ctor_set(x_82, 6, x_78); -lean_ctor_set(x_82, 7, x_79); -x_83 = lean_st_ref_set(x_51, x_82, x_61); -x_84 = lean_ctor_get(x_83, 1); -lean_inc(x_84); -lean_dec(x_83); -x_85 = lean_box(0); -x_86 = l_Lean_Environment_Replay_throwKernelException___lambda__1(x_23, x_80, x_1, x_85, x_49, x_51, x_84); -lean_dec(x_51); -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_11 = x_87; -x_12 = x_88; -goto block_22; -} +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set_tag(x_7, 1); +lean_ctor_set(x_7, 0, x_10); +return x_7; } else { -uint8_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_89 = l_Lean_Environment_Replay_throwKernelException___closed__21; -x_90 = lean_box(0); -x_91 = l_Lean_Environment_Replay_throwKernelException___lambda__1(x_23, x_89, x_1, x_90, x_49, x_51, x_55); -lean_dec(x_51); -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -lean_dec(x_91); -x_11 = x_92; -x_12 = x_93; -goto block_22; -} -} -} +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_ctor_get(x_7, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_7); +x_13 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_13, 0, x_11); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; } -LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_2); -lean_dec(x_2); -x_9 = l_Lean_Environment_Replay_throwKernelException___lambda__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_4); -return x_9; } } LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -2151,7 +1572,7 @@ static lean_object* _init_l_Lean_Environment_Replay_replayConstant___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_Environment_Replay_replayConstant___closed__1; x_2 = l_Lean_Environment_Replay_replayConstant___closed__2; -x_3 = lean_unsigned_to_nat(76u); +x_3 = lean_unsigned_to_nat(74u); x_4 = lean_unsigned_to_nat(50u); x_5 = l_Lean_Environment_Replay_replayConstant___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2221,7 +1642,7 @@ x_27 = lean_usize_sub(x_25, x_26); x_28 = lean_usize_land(x_24, x_27); x_29 = lean_array_uget(x_15, x_28); lean_dec(x_15); -x_30 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_1, x_29); +x_30 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_1, x_29); lean_dec(x_29); if (lean_obj_tag(x_30) == 0) { @@ -3941,7 +3362,7 @@ lean_inc(x_4); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); lean_dec(x_1); -x_6 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2(x_5, x_2); +x_6 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Environment_find_x3f___spec__2(x_5, x_2); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; @@ -3964,7 +3385,7 @@ x_19 = lean_usize_sub(x_17, x_18); x_20 = lean_usize_land(x_16, x_19); x_21 = lean_array_uget(x_7, x_20); lean_dec(x_7); -x_22 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_2, x_21); +x_22 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_2, x_21); lean_dec(x_21); return x_22; } @@ -4014,7 +3435,7 @@ x_39 = lean_usize_sub(x_37, x_38); x_40 = lean_usize_land(x_36, x_39); x_41 = lean_array_uget(x_27, x_40); lean_dec(x_27); -x_42 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_2, x_41); +x_42 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_2, x_41); lean_dec(x_41); return x_42; } @@ -4048,6 +3469,14 @@ static lean_object* _init_l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("Invalid constructor ", 20, 20); return x_1; } @@ -4100,7 +3529,7 @@ x_19 = lean_ctor_get(x_16, 1); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_ctor_get(x_20, 1); +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); lean_dec(x_20); x_22 = l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1(x_21, x_9); @@ -4114,7 +3543,7 @@ x_25 = l_Lean_Name_toString(x_9, x_23, x_24); x_26 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_27 = lean_string_append(x_26, x_25); lean_dec(x_25); -x_28 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_28 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_29 = lean_string_append(x_27, x_28); lean_ctor_set_tag(x_12, 18); lean_ctor_set(x_12, 0, x_29); @@ -4153,7 +3582,7 @@ x_44 = 1; x_45 = lean_usize_sub(x_43, x_44); x_46 = lean_usize_land(x_42, x_45); x_47 = lean_array_uget(x_32, x_46); -x_48 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_9, x_47); +x_48 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_9, x_47); lean_dec(x_47); if (lean_obj_tag(x_48) == 0) { @@ -4166,7 +3595,7 @@ x_51 = l_Lean_Name_toString(x_9, x_49, x_50); x_52 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_53 = lean_string_append(x_52, x_51); lean_dec(x_51); -x_54 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_54 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_55 = lean_string_append(x_53, x_54); lean_ctor_set_tag(x_30, 18); lean_ctor_set(x_30, 0, x_55); @@ -4189,7 +3618,7 @@ if (x_57 == 0) { lean_object* x_58; uint8_t x_59; x_58 = lean_ctor_get(x_56, 0); -x_59 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(x_33, x_58); +x_59 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(x_33, x_58); lean_dec(x_58); lean_dec(x_33); if (x_59 == 0) @@ -4199,10 +3628,10 @@ lean_dec(x_10); x_60 = 1; x_61 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__1; x_62 = l_Lean_Name_toString(x_9, x_60, x_61); -x_63 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; +x_63 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4; x_64 = lean_string_append(x_63, x_62); lean_dec(x_62); -x_65 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_65 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_66 = lean_string_append(x_64, x_65); lean_ctor_set_tag(x_56, 18); lean_ctor_set(x_56, 0, x_66); @@ -4229,7 +3658,7 @@ lean_object* x_69; uint8_t x_70; x_69 = lean_ctor_get(x_56, 0); lean_inc(x_69); lean_dec(x_56); -x_70 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(x_33, x_69); +x_70 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(x_33, x_69); lean_dec(x_69); lean_dec(x_33); if (x_70 == 0) @@ -4239,10 +3668,10 @@ lean_dec(x_10); x_71 = 1; x_72 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__1; x_73 = l_Lean_Name_toString(x_9, x_71, x_72); -x_74 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; +x_74 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4; x_75 = lean_string_append(x_74, x_73); lean_dec(x_73); -x_76 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_76 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_77 = lean_string_append(x_75, x_76); x_78 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_78, 0, x_77); @@ -4280,7 +3709,7 @@ x_85 = l_Lean_Name_toString(x_9, x_83, x_84); x_86 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_87 = lean_string_append(x_86, x_85); lean_dec(x_85); -x_88 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_88 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_89 = lean_string_append(x_87, x_88); lean_ctor_set_tag(x_56, 18); lean_ctor_set(x_56, 0, x_89); @@ -4298,7 +3727,7 @@ x_92 = l_Lean_Name_toString(x_9, x_90, x_91); x_93 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_94 = lean_string_append(x_93, x_92); lean_dec(x_92); -x_95 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_95 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_96 = lean_string_append(x_94, x_95); x_97 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_97, 0, x_96); @@ -4331,7 +3760,7 @@ x_110 = 1; x_111 = lean_usize_sub(x_109, x_110); x_112 = lean_usize_land(x_108, x_111); x_113 = lean_array_uget(x_98, x_112); -x_114 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_9, x_113); +x_114 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_9, x_113); lean_dec(x_113); if (lean_obj_tag(x_114) == 0) { @@ -4344,7 +3773,7 @@ x_117 = l_Lean_Name_toString(x_9, x_115, x_116); x_118 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_119 = lean_string_append(x_118, x_117); lean_dec(x_117); -x_120 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_120 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_121 = lean_string_append(x_119, x_120); x_122 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_122, 0, x_121); @@ -4370,7 +3799,7 @@ if (lean_is_exclusive(x_123)) { lean_dec_ref(x_123); x_125 = lean_box(0); } -x_126 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(x_99, x_124); +x_126 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(x_99, x_124); lean_dec(x_124); lean_dec(x_99); if (x_126 == 0) @@ -4380,10 +3809,10 @@ lean_dec(x_10); x_127 = 1; x_128 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__1; x_129 = l_Lean_Name_toString(x_9, x_127, x_128); -x_130 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; +x_130 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4; x_131 = lean_string_append(x_130, x_129); lean_dec(x_129); -x_132 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_132 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_133 = lean_string_append(x_131, x_132); if (lean_is_scalar(x_125)) { x_134 = lean_alloc_ctor(18, 1, 0); @@ -4427,7 +3856,7 @@ x_140 = l_Lean_Name_toString(x_9, x_138, x_139); x_141 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_142 = lean_string_append(x_141, x_140); lean_dec(x_140); -x_143 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_143 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_144 = lean_string_append(x_142, x_143); if (lean_is_scalar(x_137)) { x_145 = lean_alloc_ctor(18, 1, 0); @@ -4459,7 +3888,7 @@ x_150 = l_Lean_Name_toString(x_9, x_148, x_149); x_151 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_152 = lean_string_append(x_151, x_150); lean_dec(x_150); -x_153 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_153 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_154 = lean_string_append(x_152, x_153); lean_ctor_set_tag(x_30, 18); lean_ctor_set(x_30, 0, x_154); @@ -4477,7 +3906,7 @@ x_157 = l_Lean_Name_toString(x_9, x_155, x_156); x_158 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_159 = lean_string_append(x_158, x_157); lean_dec(x_157); -x_160 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_160 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_161 = lean_string_append(x_159, x_160); x_162 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_162, 0, x_161); @@ -4499,7 +3928,7 @@ lean_dec(x_16); x_165 = lean_ctor_get(x_163, 0); lean_inc(x_165); lean_dec(x_163); -x_166 = lean_ctor_get(x_165, 1); +x_166 = lean_ctor_get(x_165, 0); lean_inc(x_166); lean_dec(x_165); x_167 = l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1(x_166, x_9); @@ -4513,7 +3942,7 @@ x_170 = l_Lean_Name_toString(x_9, x_168, x_169); x_171 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_172 = lean_string_append(x_171, x_170); lean_dec(x_170); -x_173 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_173 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_174 = lean_string_append(x_172, x_173); lean_ctor_set_tag(x_12, 18); lean_ctor_set(x_12, 0, x_174); @@ -4557,7 +3986,7 @@ x_190 = 1; x_191 = lean_usize_sub(x_189, x_190); x_192 = lean_usize_land(x_188, x_191); x_193 = lean_array_uget(x_177, x_192); -x_194 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_9, x_193); +x_194 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_9, x_193); lean_dec(x_193); if (lean_obj_tag(x_194) == 0) { @@ -4570,7 +3999,7 @@ x_197 = l_Lean_Name_toString(x_9, x_195, x_196); x_198 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_199 = lean_string_append(x_198, x_197); lean_dec(x_197); -x_200 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_200 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_201 = lean_string_append(x_199, x_200); if (lean_is_scalar(x_179)) { x_202 = lean_alloc_ctor(18, 1, 0); @@ -4603,7 +4032,7 @@ if (lean_is_exclusive(x_204)) { lean_dec_ref(x_204); x_206 = lean_box(0); } -x_207 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(x_178, x_205); +x_207 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(x_178, x_205); lean_dec(x_205); lean_dec(x_178); if (x_207 == 0) @@ -4613,10 +4042,10 @@ lean_dec(x_10); x_208 = 1; x_209 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__1; x_210 = l_Lean_Name_toString(x_9, x_208, x_209); -x_211 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; +x_211 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4; x_212 = lean_string_append(x_211, x_210); lean_dec(x_210); -x_213 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_213 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_214 = lean_string_append(x_212, x_213); if (lean_is_scalar(x_206)) { x_215 = lean_alloc_ctor(18, 1, 0); @@ -4660,7 +4089,7 @@ x_222 = l_Lean_Name_toString(x_9, x_220, x_221); x_223 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_224 = lean_string_append(x_223, x_222); lean_dec(x_222); -x_225 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_225 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_226 = lean_string_append(x_224, x_225); if (lean_is_scalar(x_219)) { x_227 = lean_alloc_ctor(18, 1, 0); @@ -4693,7 +4122,7 @@ x_232 = l_Lean_Name_toString(x_9, x_230, x_231); x_233 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_234 = lean_string_append(x_233, x_232); lean_dec(x_232); -x_235 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_235 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_236 = lean_string_append(x_234, x_235); if (lean_is_scalar(x_229)) { x_237 = lean_alloc_ctor(18, 1, 0); @@ -4733,7 +4162,7 @@ if (lean_is_exclusive(x_240)) { x_244 = lean_ctor_get(x_241, 0); lean_inc(x_244); lean_dec(x_241); -x_245 = lean_ctor_get(x_244, 1); +x_245 = lean_ctor_get(x_244, 0); lean_inc(x_245); lean_dec(x_244); x_246 = l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1(x_245, x_9); @@ -4747,7 +4176,7 @@ x_249 = l_Lean_Name_toString(x_9, x_247, x_248); x_250 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_251 = lean_string_append(x_250, x_249); lean_dec(x_249); -x_252 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_252 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_253 = lean_string_append(x_251, x_252); x_254 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_254, 0, x_253); @@ -4795,7 +4224,7 @@ x_270 = 1; x_271 = lean_usize_sub(x_269, x_270); x_272 = lean_usize_land(x_268, x_271); x_273 = lean_array_uget(x_257, x_272); -x_274 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_9, x_273); +x_274 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_9, x_273); lean_dec(x_273); if (lean_obj_tag(x_274) == 0) { @@ -4808,7 +4237,7 @@ x_277 = l_Lean_Name_toString(x_9, x_275, x_276); x_278 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_279 = lean_string_append(x_278, x_277); lean_dec(x_277); -x_280 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_280 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_281 = lean_string_append(x_279, x_280); if (lean_is_scalar(x_259)) { x_282 = lean_alloc_ctor(18, 1, 0); @@ -4846,7 +4275,7 @@ if (lean_is_exclusive(x_284)) { lean_dec_ref(x_284); x_286 = lean_box(0); } -x_287 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2774_(x_258, x_285); +x_287 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2900_(x_258, x_285); lean_dec(x_285); lean_dec(x_258); if (x_287 == 0) @@ -4856,10 +4285,10 @@ lean_dec(x_10); x_288 = 1; x_289 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__1; x_290 = l_Lean_Name_toString(x_9, x_288, x_289); -x_291 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; +x_291 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4; x_292 = lean_string_append(x_291, x_290); lean_dec(x_290); -x_293 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_293 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_294 = lean_string_append(x_292, x_293); if (lean_is_scalar(x_286)) { x_295 = lean_alloc_ctor(18, 1, 0); @@ -4909,7 +4338,7 @@ x_302 = l_Lean_Name_toString(x_9, x_300, x_301); x_303 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_304 = lean_string_append(x_303, x_302); lean_dec(x_302); -x_305 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_305 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_306 = lean_string_append(x_304, x_305); if (lean_is_scalar(x_299)) { x_307 = lean_alloc_ctor(18, 1, 0); @@ -4947,7 +4376,7 @@ x_312 = l_Lean_Name_toString(x_9, x_310, x_311); x_313 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2; x_314 = lean_string_append(x_313, x_312); lean_dec(x_312); -x_315 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_315 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_316 = lean_string_append(x_314, x_315); if (lean_is_scalar(x_309)) { x_317 = lean_alloc_ctor(18, 1, 0); @@ -5162,7 +4591,7 @@ x_19 = lean_ctor_get(x_16, 1); x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); -x_21 = lean_ctor_get(x_20, 1); +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); lean_dec(x_20); x_22 = l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1(x_21, x_9); @@ -5176,7 +4605,7 @@ x_25 = l_Lean_Name_toString(x_9, x_23, x_24); x_26 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_27 = lean_string_append(x_26, x_25); lean_dec(x_25); -x_28 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_28 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_29 = lean_string_append(x_27, x_28); lean_ctor_set_tag(x_12, 18); lean_ctor_set(x_12, 0, x_29); @@ -5215,7 +4644,7 @@ x_44 = 1; x_45 = lean_usize_sub(x_43, x_44); x_46 = lean_usize_land(x_42, x_45); x_47 = lean_array_uget(x_32, x_46); -x_48 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_9, x_47); +x_48 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_9, x_47); lean_dec(x_47); if (lean_obj_tag(x_48) == 0) { @@ -5228,7 +4657,7 @@ x_51 = l_Lean_Name_toString(x_9, x_49, x_50); x_52 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_53 = lean_string_append(x_52, x_51); lean_dec(x_51); -x_54 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_54 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_55 = lean_string_append(x_53, x_54); lean_ctor_set_tag(x_30, 18); lean_ctor_set(x_30, 0, x_55); @@ -5251,7 +4680,7 @@ if (x_57 == 0) { lean_object* x_58; uint8_t x_59; x_58 = lean_ctor_get(x_56, 0); -x_59 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(x_33, x_58); +x_59 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(x_33, x_58); lean_dec(x_58); lean_dec(x_33); if (x_59 == 0) @@ -5264,7 +4693,7 @@ x_62 = l_Lean_Name_toString(x_9, x_60, x_61); x_63 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2; x_64 = lean_string_append(x_63, x_62); lean_dec(x_62); -x_65 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_65 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_66 = lean_string_append(x_64, x_65); lean_ctor_set_tag(x_56, 18); lean_ctor_set(x_56, 0, x_66); @@ -5291,7 +4720,7 @@ lean_object* x_69; uint8_t x_70; x_69 = lean_ctor_get(x_56, 0); lean_inc(x_69); lean_dec(x_56); -x_70 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(x_33, x_69); +x_70 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(x_33, x_69); lean_dec(x_69); lean_dec(x_33); if (x_70 == 0) @@ -5304,7 +4733,7 @@ x_73 = l_Lean_Name_toString(x_9, x_71, x_72); x_74 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2; x_75 = lean_string_append(x_74, x_73); lean_dec(x_73); -x_76 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_76 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_77 = lean_string_append(x_75, x_76); x_78 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_78, 0, x_77); @@ -5342,7 +4771,7 @@ x_85 = l_Lean_Name_toString(x_9, x_83, x_84); x_86 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_87 = lean_string_append(x_86, x_85); lean_dec(x_85); -x_88 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_88 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_89 = lean_string_append(x_87, x_88); lean_ctor_set_tag(x_56, 18); lean_ctor_set(x_56, 0, x_89); @@ -5360,7 +4789,7 @@ x_92 = l_Lean_Name_toString(x_9, x_90, x_91); x_93 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_94 = lean_string_append(x_93, x_92); lean_dec(x_92); -x_95 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_95 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_96 = lean_string_append(x_94, x_95); x_97 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_97, 0, x_96); @@ -5393,7 +4822,7 @@ x_110 = 1; x_111 = lean_usize_sub(x_109, x_110); x_112 = lean_usize_land(x_108, x_111); x_113 = lean_array_uget(x_98, x_112); -x_114 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_9, x_113); +x_114 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_9, x_113); lean_dec(x_113); if (lean_obj_tag(x_114) == 0) { @@ -5406,7 +4835,7 @@ x_117 = l_Lean_Name_toString(x_9, x_115, x_116); x_118 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_119 = lean_string_append(x_118, x_117); lean_dec(x_117); -x_120 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_120 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_121 = lean_string_append(x_119, x_120); x_122 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_122, 0, x_121); @@ -5432,7 +4861,7 @@ if (lean_is_exclusive(x_123)) { lean_dec_ref(x_123); x_125 = lean_box(0); } -x_126 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(x_99, x_124); +x_126 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(x_99, x_124); lean_dec(x_124); lean_dec(x_99); if (x_126 == 0) @@ -5445,7 +4874,7 @@ x_129 = l_Lean_Name_toString(x_9, x_127, x_128); x_130 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2; x_131 = lean_string_append(x_130, x_129); lean_dec(x_129); -x_132 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_132 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_133 = lean_string_append(x_131, x_132); if (lean_is_scalar(x_125)) { x_134 = lean_alloc_ctor(18, 1, 0); @@ -5489,7 +4918,7 @@ x_140 = l_Lean_Name_toString(x_9, x_138, x_139); x_141 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_142 = lean_string_append(x_141, x_140); lean_dec(x_140); -x_143 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_143 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_144 = lean_string_append(x_142, x_143); if (lean_is_scalar(x_137)) { x_145 = lean_alloc_ctor(18, 1, 0); @@ -5521,7 +4950,7 @@ x_150 = l_Lean_Name_toString(x_9, x_148, x_149); x_151 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_152 = lean_string_append(x_151, x_150); lean_dec(x_150); -x_153 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_153 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_154 = lean_string_append(x_152, x_153); lean_ctor_set_tag(x_30, 18); lean_ctor_set(x_30, 0, x_154); @@ -5539,7 +4968,7 @@ x_157 = l_Lean_Name_toString(x_9, x_155, x_156); x_158 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_159 = lean_string_append(x_158, x_157); lean_dec(x_157); -x_160 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_160 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_161 = lean_string_append(x_159, x_160); x_162 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_162, 0, x_161); @@ -5561,7 +4990,7 @@ lean_dec(x_16); x_165 = lean_ctor_get(x_163, 0); lean_inc(x_165); lean_dec(x_163); -x_166 = lean_ctor_get(x_165, 1); +x_166 = lean_ctor_get(x_165, 0); lean_inc(x_166); lean_dec(x_165); x_167 = l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1(x_166, x_9); @@ -5575,7 +5004,7 @@ x_170 = l_Lean_Name_toString(x_9, x_168, x_169); x_171 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_172 = lean_string_append(x_171, x_170); lean_dec(x_170); -x_173 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_173 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_174 = lean_string_append(x_172, x_173); lean_ctor_set_tag(x_12, 18); lean_ctor_set(x_12, 0, x_174); @@ -5619,7 +5048,7 @@ x_190 = 1; x_191 = lean_usize_sub(x_189, x_190); x_192 = lean_usize_land(x_188, x_191); x_193 = lean_array_uget(x_177, x_192); -x_194 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_9, x_193); +x_194 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_9, x_193); lean_dec(x_193); if (lean_obj_tag(x_194) == 0) { @@ -5632,7 +5061,7 @@ x_197 = l_Lean_Name_toString(x_9, x_195, x_196); x_198 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_199 = lean_string_append(x_198, x_197); lean_dec(x_197); -x_200 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_200 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_201 = lean_string_append(x_199, x_200); if (lean_is_scalar(x_179)) { x_202 = lean_alloc_ctor(18, 1, 0); @@ -5665,7 +5094,7 @@ if (lean_is_exclusive(x_204)) { lean_dec_ref(x_204); x_206 = lean_box(0); } -x_207 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(x_178, x_205); +x_207 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(x_178, x_205); lean_dec(x_205); lean_dec(x_178); if (x_207 == 0) @@ -5678,7 +5107,7 @@ x_210 = l_Lean_Name_toString(x_9, x_208, x_209); x_211 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2; x_212 = lean_string_append(x_211, x_210); lean_dec(x_210); -x_213 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_213 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_214 = lean_string_append(x_212, x_213); if (lean_is_scalar(x_206)) { x_215 = lean_alloc_ctor(18, 1, 0); @@ -5722,7 +5151,7 @@ x_222 = l_Lean_Name_toString(x_9, x_220, x_221); x_223 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_224 = lean_string_append(x_223, x_222); lean_dec(x_222); -x_225 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_225 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_226 = lean_string_append(x_224, x_225); if (lean_is_scalar(x_219)) { x_227 = lean_alloc_ctor(18, 1, 0); @@ -5755,7 +5184,7 @@ x_232 = l_Lean_Name_toString(x_9, x_230, x_231); x_233 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_234 = lean_string_append(x_233, x_232); lean_dec(x_232); -x_235 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_235 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_236 = lean_string_append(x_234, x_235); if (lean_is_scalar(x_229)) { x_237 = lean_alloc_ctor(18, 1, 0); @@ -5795,7 +5224,7 @@ if (lean_is_exclusive(x_240)) { x_244 = lean_ctor_get(x_241, 0); lean_inc(x_244); lean_dec(x_241); -x_245 = lean_ctor_get(x_244, 1); +x_245 = lean_ctor_get(x_244, 0); lean_inc(x_245); lean_dec(x_244); x_246 = l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1(x_245, x_9); @@ -5809,7 +5238,7 @@ x_249 = l_Lean_Name_toString(x_9, x_247, x_248); x_250 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_251 = lean_string_append(x_250, x_249); lean_dec(x_249); -x_252 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_252 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_253 = lean_string_append(x_251, x_252); x_254 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_254, 0, x_253); @@ -5857,7 +5286,7 @@ x_270 = 1; x_271 = lean_usize_sub(x_269, x_270); x_272 = lean_usize_land(x_268, x_271); x_273 = lean_array_uget(x_257, x_272); -x_274 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Environment_find_x3f___spec__5(x_9, x_273); +x_274 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Kernel_Environment_find_x3f___spec__5(x_9, x_273); lean_dec(x_273); if (lean_obj_tag(x_274) == 0) { @@ -5870,7 +5299,7 @@ x_277 = l_Lean_Name_toString(x_9, x_275, x_276); x_278 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_279 = lean_string_append(x_278, x_277); lean_dec(x_277); -x_280 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_280 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_281 = lean_string_append(x_279, x_280); if (lean_is_scalar(x_259)) { x_282 = lean_alloc_ctor(18, 1, 0); @@ -5908,7 +5337,7 @@ if (lean_is_exclusive(x_284)) { lean_dec_ref(x_284); x_286 = lean_box(0); } -x_287 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3206_(x_258, x_285); +x_287 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_3332_(x_258, x_285); lean_dec(x_285); lean_dec(x_258); if (x_287 == 0) @@ -5921,7 +5350,7 @@ x_290 = l_Lean_Name_toString(x_9, x_288, x_289); x_291 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2; x_292 = lean_string_append(x_291, x_290); lean_dec(x_290); -x_293 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_293 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_294 = lean_string_append(x_292, x_293); if (lean_is_scalar(x_286)) { x_295 = lean_alloc_ctor(18, 1, 0); @@ -5971,7 +5400,7 @@ x_302 = l_Lean_Name_toString(x_9, x_300, x_301); x_303 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_304 = lean_string_append(x_303, x_302); lean_dec(x_302); -x_305 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_305 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_306 = lean_string_append(x_304, x_305); if (lean_is_scalar(x_299)) { x_307 = lean_alloc_ctor(18, 1, 0); @@ -6009,7 +5438,7 @@ x_312 = l_Lean_Name_toString(x_9, x_310, x_311); x_313 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1; x_314 = lean_string_append(x_313, x_312); lean_dec(x_312); -x_315 = l_Lean_Environment_Replay_throwKernelException___closed__3; +x_315 = l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3; x_316 = lean_string_append(x_314, x_315); if (lean_is_scalar(x_309)) { x_317 = lean_alloc_ctor(18, 1, 0); @@ -6493,49 +5922,6 @@ lean_dec_ref(res); res = initialize_Lean_Util_FoldConsts(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1 = _init_l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1); -l_Lean_Environment_Replay_throwKernelException___closed__1 = _init_l_Lean_Environment_Replay_throwKernelException___closed__1(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__1); -l_Lean_Environment_Replay_throwKernelException___closed__2 = _init_l_Lean_Environment_Replay_throwKernelException___closed__2(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__2); -l_Lean_Environment_Replay_throwKernelException___closed__3 = _init_l_Lean_Environment_Replay_throwKernelException___closed__3(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__3); -l_Lean_Environment_Replay_throwKernelException___closed__4 = _init_l_Lean_Environment_Replay_throwKernelException___closed__4(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__4); -l_Lean_Environment_Replay_throwKernelException___closed__5 = _init_l_Lean_Environment_Replay_throwKernelException___closed__5(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__5); -l_Lean_Environment_Replay_throwKernelException___closed__6 = _init_l_Lean_Environment_Replay_throwKernelException___closed__6(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__6); -l_Lean_Environment_Replay_throwKernelException___closed__7 = _init_l_Lean_Environment_Replay_throwKernelException___closed__7(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__7); -l_Lean_Environment_Replay_throwKernelException___closed__8 = _init_l_Lean_Environment_Replay_throwKernelException___closed__8(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__8); -l_Lean_Environment_Replay_throwKernelException___closed__9 = _init_l_Lean_Environment_Replay_throwKernelException___closed__9(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__9); -l_Lean_Environment_Replay_throwKernelException___closed__10 = _init_l_Lean_Environment_Replay_throwKernelException___closed__10(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__10); -l_Lean_Environment_Replay_throwKernelException___closed__11 = _init_l_Lean_Environment_Replay_throwKernelException___closed__11(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__11); -l_Lean_Environment_Replay_throwKernelException___closed__12 = _init_l_Lean_Environment_Replay_throwKernelException___closed__12(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__12); -l_Lean_Environment_Replay_throwKernelException___closed__13 = _init_l_Lean_Environment_Replay_throwKernelException___closed__13(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__13); -l_Lean_Environment_Replay_throwKernelException___closed__14 = _init_l_Lean_Environment_Replay_throwKernelException___closed__14(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__14); -l_Lean_Environment_Replay_throwKernelException___closed__15 = _init_l_Lean_Environment_Replay_throwKernelException___closed__15(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__15); -l_Lean_Environment_Replay_throwKernelException___closed__16 = _init_l_Lean_Environment_Replay_throwKernelException___closed__16(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__16); -l_Lean_Environment_Replay_throwKernelException___closed__17 = _init_l_Lean_Environment_Replay_throwKernelException___closed__17(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__17); -l_Lean_Environment_Replay_throwKernelException___closed__18 = _init_l_Lean_Environment_Replay_throwKernelException___closed__18(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__18); -l_Lean_Environment_Replay_throwKernelException___closed__19 = _init_l_Lean_Environment_Replay_throwKernelException___closed__19(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__19); -l_Lean_Environment_Replay_throwKernelException___closed__20 = _init_l_Lean_Environment_Replay_throwKernelException___closed__20(); -lean_mark_persistent(l_Lean_Environment_Replay_throwKernelException___closed__20); -l_Lean_Environment_Replay_throwKernelException___closed__21 = _init_l_Lean_Environment_Replay_throwKernelException___closed__21(); l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__1 = _init_l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__1(); lean_mark_persistent(l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__1); l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__2 = _init_l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__2(); @@ -6566,6 +5952,8 @@ l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructor lean_mark_persistent(l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__2); l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3 = _init_l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3(); lean_mark_persistent(l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__3); +l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4 = _init_l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___closed__4); l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1 = _init_l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1(); lean_mark_persistent(l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__1); l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2 = _init_l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2(); diff --git a/stage0/stdlib/Lean/ResolveName.c b/stage0/stdlib/Lean/ResolveName.c index 7e3430dff07e..8059375b155c 100644 --- a/stage0/stdlib/Lean/ResolveName.c +++ b/stage0/stdlib/Lean/ResolveName.c @@ -3227,7 +3227,6 @@ x_7 = l___private_Lean_ResolveName_0__Lean_ResolveName_containsDeclOrReserved(x_ if (x_7 == 0) { lean_object* x_8; uint8_t x_9; -lean_inc(x_1); x_8 = l_Lean_mkPrivateName(x_1, x_4); lean_inc(x_8); x_9 = l___private_Lean_ResolveName_0__Lean_ResolveName_containsDeclOrReserved(x_1, x_8); @@ -3273,7 +3272,6 @@ return x_14; else { lean_object* x_15; uint8_t x_16; -lean_inc(x_1); x_15 = l_Lean_mkPrivateName(x_1, x_4); lean_inc(x_15); x_16 = l___private_Lean_ResolveName_0__Lean_ResolveName_containsDeclOrReserved(x_1, x_15); @@ -3347,7 +3345,6 @@ x_7 = l___private_Lean_ResolveName_0__Lean_ResolveName_containsDeclOrReserved(x_ if (x_7 == 0) { lean_object* x_8; uint8_t x_9; -lean_inc(x_1); x_8 = l_Lean_mkPrivateName(x_1, x_6); lean_inc(x_8); x_9 = l___private_Lean_ResolveName_0__Lean_ResolveName_containsDeclOrReserved(x_1, x_8); @@ -3958,7 +3955,6 @@ lean_inc(x_13); lean_inc(x_1); x_16 = l___private_Lean_ResolveName_0__Lean_ResolveName_containsDeclOrReserved(x_1, x_13); lean_inc(x_13); -lean_inc(x_1); x_17 = l_Lean_mkPrivateName(x_1, x_13); lean_inc(x_17); lean_inc(x_1); diff --git a/stage0/stdlib/Lean/Server/CodeActions/Basic.c b/stage0/stdlib/Lean/Server/CodeActions/Basic.c index 0f8b24763c5e..d628a224785e 100644 --- a/stage0/stdlib/Lean/Server/CodeActions/Basic.c +++ b/stage0/stdlib/Lean/Server/CodeActions/Basic.c @@ -41,6 +41,7 @@ lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_CodeAction_getFileSource_x21___closed__2; lean_object* l_Lean_ConstantInfo_type(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_fromJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_125____closed__11; LEAN_EXPORT lean_object* l_Lean_Server_instFromJsonCodeActionResolveData; static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____closed__20; @@ -92,6 +93,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_handleCodeActionResolve(lean_object*, lea LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at_Lean_Server_handleCodeAction___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_458____closed__2; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_withNewLocals___spec__1(lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_instCodeActionProviderInhabited___closed__1; static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_458_(lean_object*); @@ -145,14 +147,12 @@ static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___ lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_instCodeActionProviderInhabited___lambda__1(lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_fromJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_125____closed__10; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_handleCodeAction(lean_object*, lean_object*, lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_517____closed__15; static lean_object* l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_fromJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_125____closed__7; LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_412_(lean_object*); @@ -423,7 +423,7 @@ x_24 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_24, 0, x_7); lean_ctor_set(x_24, 1, x_23); x_25 = l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_toJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_59____closed__5; -x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_24, x_25); +x_26 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_24, x_25); x_27 = l_Lean_Json_mkObj(x_26); return x_27; } @@ -740,7 +740,7 @@ x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Server_CodeActions_Basic_0__Lean_Server_toJsonCodeActionResolveData____x40_Lean_Server_CodeActions_Basic___hyg_59____closed__4; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_273____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; diff --git a/stage0/stdlib/Lean/Server/Completion.c b/stage0/stdlib/Lean/Server/Completion.c index db1af91e4c1d..1418b57b0551 100644 --- a/stage0/stdlib/Lean/Server/Completion.c +++ b/stage0/stdlib/Lean/Server/Completion.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Server.Completion -// Imports: Lean.Server.Completion.CompletionCollectors +// Imports: Lean.Server.Completion.CompletionCollectors Std.Data.HashMap #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -42,7 +42,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Co LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___lambda__1___boxed(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__20; -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2; @@ -130,6 +129,7 @@ static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__ lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__17(size_t, size_t, lean_object*); static uint64_t l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1; +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(lean_object*); extern lean_object* l_Lean_Lsp_instHashableCompletionItemKind; static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__27; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -843,7 +843,7 @@ x_40 = lean_ctor_get(x_34, 0); lean_inc(x_40); lean_dec(x_34); x_41 = 11; -x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_40); +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_40); lean_dec(x_40); x_43 = 13; x_44 = lean_uint64_mix_hash(x_42, x_43); @@ -888,7 +888,7 @@ lean_object* x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; uint64_t x_65; u x_61 = lean_ctor_get(x_34, 0); lean_inc(x_61); lean_dec(x_34); -x_62 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_61); +x_62 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_61); lean_dec(x_61); x_63 = lean_uint64_mix_hash(x_62, x_55); x_64 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; @@ -926,7 +926,7 @@ lean_object* x_76; uint64_t x_77; uint64_t x_78; uint64_t x_79; uint64_t x_80; u x_76 = lean_ctor_get(x_34, 0); lean_inc(x_76); lean_dec(x_34); -x_77 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_76); +x_77 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_76); lean_dec(x_76); x_78 = lean_uint64_mix_hash(x_77, x_55); x_79 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; @@ -966,7 +966,7 @@ lean_object* x_95; uint64_t x_96; uint64_t x_97; uint64_t x_98; uint64_t x_99; u x_95 = lean_ctor_get(x_34, 0); lean_inc(x_95); lean_dec(x_34); -x_96 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_95); +x_96 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_95); lean_dec(x_95); x_97 = lean_uint64_mix_hash(x_96, x_55); x_98 = lean_uint64_mix_hash(x_88, x_97); @@ -2193,7 +2193,7 @@ x_146 = lean_ctor_get(x_140, 0); lean_inc(x_146); lean_dec(x_140); x_147 = 11; -x_148 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_146); +x_148 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_146); lean_dec(x_146); x_149 = 13; x_150 = lean_uint64_mix_hash(x_148, x_149); @@ -2238,7 +2238,7 @@ lean_object* x_167; uint64_t x_168; uint64_t x_169; uint64_t x_170; uint64_t x_1 x_167 = lean_ctor_get(x_140, 0); lean_inc(x_167); lean_dec(x_140); -x_168 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_167); +x_168 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_167); lean_dec(x_167); x_169 = lean_uint64_mix_hash(x_168, x_161); x_170 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; @@ -2276,7 +2276,7 @@ lean_object* x_182; uint64_t x_183; uint64_t x_184; uint64_t x_185; uint64_t x_1 x_182 = lean_ctor_get(x_140, 0); lean_inc(x_182); lean_dec(x_140); -x_183 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_182); +x_183 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_182); lean_dec(x_182); x_184 = lean_uint64_mix_hash(x_183, x_161); x_185 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; @@ -2316,7 +2316,7 @@ lean_object* x_201; uint64_t x_202; uint64_t x_203; uint64_t x_204; uint64_t x_2 x_201 = lean_ctor_get(x_140, 0); lean_inc(x_201); lean_dec(x_140); -x_202 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6462_(x_201); +x_202 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6277_(x_201); lean_dec(x_201); x_203 = lean_uint64_mix_hash(x_202, x_161); x_204 = lean_uint64_mix_hash(x_194, x_203); @@ -4566,6 +4566,7 @@ return x_13; } } lean_object* initialize_Lean_Server_Completion_CompletionCollectors(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Data_HashMap(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Server_Completion(uint8_t builtin, lean_object* w) { lean_object * res; @@ -4574,6 +4575,9 @@ _G_initialized = true; res = initialize_Lean_Server_Completion_CompletionCollectors(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Std_Data_HashMap(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1(); l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2(); l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3(); diff --git a/stage0/stdlib/Lean/Server/Completion/CompletionCollectors.c b/stage0/stdlib/Lean/Server/Completion/CompletionCollectors.c index 8ebb687331cb..864d5f91d299 100644 --- a/stage0/stdlib/Lean/Server/Completion/CompletionCollectors.c +++ b/stage0/stdlib/Lean/Server/Completion/CompletionCollectors.c @@ -124,7 +124,7 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___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_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___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_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___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* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1; @@ -1396,8 +1396,7 @@ x_12 = lean_ctor_get(x_9, 1); x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); -lean_inc(x_1); -x_14 = lean_environment_find(x_13, x_1); +x_14 = l_Lean_Environment_find_x3f(x_13, x_1); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -1437,8 +1436,7 @@ lean_dec(x_9); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -lean_inc(x_1); -x_26 = lean_environment_find(x_25, x_1); +x_26 = l_Lean_Environment_find_x3f(x_25, x_1); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; @@ -2801,8 +2799,7 @@ x_14 = lean_ctor_get(x_11, 1); x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); lean_dec(x_13); -lean_inc(x_2); -x_16 = lean_environment_find(x_15, x_2); +x_16 = l_Lean_Environment_find_x3f(x_15, x_2); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; @@ -2955,8 +2952,7 @@ lean_dec(x_11); x_42 = lean_ctor_get(x_40, 0); lean_inc(x_42); lean_dec(x_40); -lean_inc(x_2); -x_43 = lean_environment_find(x_42, x_2); +x_43 = l_Lean_Environment_find_x3f(x_42, x_2); if (lean_obj_tag(x_43) == 0) { lean_object* x_44; lean_object* x_45; @@ -3318,6 +3314,7 @@ lean_inc(x_15); lean_dec(x_14); lean_inc(x_11); x_16 = l_Lean_mkPrivateName(x_15, x_11); +lean_dec(x_15); x_17 = lean_name_eq(x_16, x_1); lean_dec(x_1); lean_dec(x_16); @@ -3349,6 +3346,7 @@ lean_inc(x_21); lean_dec(x_19); lean_inc(x_11); x_22 = l_Lean_mkPrivateName(x_21, x_11); +lean_dec(x_21); x_23 = lean_name_eq(x_22, x_1); lean_dec(x_1); lean_dec(x_22); @@ -3397,6 +3395,7 @@ lean_inc(x_32); lean_dec(x_29); lean_inc(x_27); x_33 = l_Lean_mkPrivateName(x_32, x_27); +lean_dec(x_32); x_34 = lean_name_eq(x_33, x_1); lean_dec(x_1); lean_dec(x_33); @@ -15443,7 +15442,7 @@ if (x_19 == 0) lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_dec(x_17); lean_dec(x_16); -x_20 = lean_ctor_get(x_12, 1); +x_20 = lean_ctor_get(x_12, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); @@ -15464,7 +15463,7 @@ if (x_24 == 0) lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_17); lean_dec(x_16); -x_25 = lean_ctor_get(x_12, 1); +x_25 = lean_ctor_get(x_12, 0); lean_inc(x_25); x_26 = lean_ctor_get(x_25, 1); lean_inc(x_26); @@ -15498,7 +15497,7 @@ lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); lean_dec(x_32); -x_34 = lean_ctor_get(x_12, 1); +x_34 = lean_ctor_get(x_12, 0); lean_inc(x_34); x_35 = lean_ctor_get(x_34, 1); lean_inc(x_35); diff --git a/stage0/stdlib/Lean/Server/Completion/CompletionResolution.c b/stage0/stdlib/Lean/Server/Completion/CompletionResolution.c index c7faa7c7294e..f64e4b80afb2 100644 --- a/stage0/stdlib/Lean/Server/Completion/CompletionResolution.c +++ b/stage0/stdlib/Lean/Server/Completion/CompletionResolution.c @@ -31,7 +31,7 @@ uint8_t l_Lean_Name_isAnonymous(lean_object*); static lean_object* l_Lean_Lsp_instToJsonResolvableCompletionItemData___closed__1; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonResolvableCompletionItemData___closed__1; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_262____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_24____lambda__3___closed__1; static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_262____closed__15; @@ -1785,7 +1785,8 @@ lean_dec(x_12); x_77 = lean_ctor_get(x_2, 0); lean_inc(x_77); lean_dec(x_2); -x_78 = lean_environment_find(x_11, x_77); +x_78 = l_Lean_Environment_find_x3f(x_11, x_77); +lean_dec(x_77); if (lean_obj_tag(x_78) == 0) { lean_object* x_79; diff --git a/stage0/stdlib/Lean/Server/Completion/EligibleHeaderDecls.c b/stage0/stdlib/Lean/Server/Completion/EligibleHeaderDecls.c index 334222cfdbe5..91bca48108ec 100644 --- a/stage0/stdlib/Lean/Server/Completion/EligibleHeaderDecls.c +++ b/stage0/stdlib/Lean/Server/Completion/EligibleHeaderDecls.c @@ -24,11 +24,11 @@ lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Server_Completion_allowCompletion___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_eligibleHeaderDeclsRef; +lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -40,6 +40,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completi LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(lean_object*, lean_object*); +uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls(lean_object*, lean_object*); static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -57,7 +58,7 @@ uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM(lean_object*); static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; uint64_t lean_uint64_xor(uint64_t, uint64_t); -lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); size_t lean_usize_add(size_t, size_t); @@ -66,7 +67,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg(lean_ lean_object* lean_array_uget(lean_object*, size_t); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(lean_object*); static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3; lean_object* lean_array_get_size(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -163,7 +163,7 @@ x_26 = 1; x_27 = lean_usize_sub(x_25, x_26); x_28 = lean_usize_land(x_24, x_27); x_29 = lean_array_uget(x_15, x_28); -x_30 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_7, x_29); +x_30 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_7, x_29); if (x_30 == 0) { lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; @@ -184,7 +184,7 @@ lean_dec(x_37); if (x_39 == 0) { lean_object* x_40; lean_object* x_41; -x_40 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_33); +x_40 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_33); lean_ctor_set(x_4, 1, x_40); lean_ctor_set(x_4, 0, x_32); x_41 = lean_box(0); @@ -209,7 +209,7 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean lean_free_object(x_3); x_45 = lean_box(0); x_46 = lean_array_uset(x_15, x_28, x_45); -x_47 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_7, x_8, x_29); +x_47 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(x_7, x_8, x_29); x_48 = lean_array_uset(x_46, x_28, x_47); lean_ctor_set(x_4, 1, x_48); x_49 = lean_box(0); @@ -241,7 +241,7 @@ x_63 = 1; x_64 = lean_usize_sub(x_62, x_63); x_65 = lean_usize_land(x_61, x_64); x_66 = lean_array_uget(x_52, x_65); -x_67 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_7, x_66); +x_67 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_7, x_66); if (x_67 == 0) { lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; @@ -262,7 +262,7 @@ lean_dec(x_74); if (x_76 == 0) { lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_70); +x_77 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_70); x_78 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_78, 0, x_69); lean_ctor_set(x_78, 1, x_77); @@ -291,7 +291,7 @@ lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean lean_free_object(x_3); x_84 = lean_box(0); x_85 = lean_array_uset(x_52, x_65, x_84); -x_86 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_7, x_8, x_66); +x_86 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(x_7, x_8, x_66); x_87 = lean_array_uset(x_85, x_65, x_86); x_88 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_88, 0, x_51); @@ -358,7 +358,7 @@ x_110 = 1; x_111 = lean_usize_sub(x_109, x_110); x_112 = lean_usize_land(x_108, x_111); x_113 = lean_array_uget(x_98, x_112); -x_114 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_91, x_113); +x_114 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_91, x_113); if (x_114 == 0) { lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; @@ -382,7 +382,7 @@ lean_dec(x_122); if (x_124 == 0) { lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_118); +x_125 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__7(x_118); if (lean_is_scalar(x_99)) { x_126 = lean_alloc_ctor(0, 2, 0); } else { @@ -418,7 +418,7 @@ else lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; x_132 = lean_box(0); x_133 = lean_array_uset(x_98, x_112, x_132); -x_134 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_91, x_92, x_113); +x_134 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__10(x_91, x_92, x_113); x_135 = lean_array_uset(x_133, x_112, x_134); if (lean_is_scalar(x_99)) { x_136 = lean_alloc_ctor(0, 2, 0); @@ -538,7 +538,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = lean_ctor_get(x_4, 1); lean_inc(x_6); lean_dec(x_4); -x_7 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_1, 0); lean_inc(x_7); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); @@ -930,7 +930,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lam _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 0); lean_inc(x_5); x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); @@ -1128,12 +1128,12 @@ x_15 = 1; x_16 = lean_usize_sub(x_14, x_15); x_17 = lean_usize_land(x_13, x_16); x_18 = lean_array_uget(x_4, x_17); -x_19 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_3, x_18); +x_19 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Kernel_Environment_add___spec__6(x_3, x_18); lean_dec(x_18); if (x_19 == 0) { lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_2, 1); +x_20 = lean_ctor_get(x_2, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); diff --git a/stage0/stdlib/Lean/Server/FileWorker.c b/stage0/stdlib/Lean/Server/FileWorker.c index ee99f83736a4..46a1a15c7791 100644 --- a/stage0/stdlib/Lean/Server/FileWorker.c +++ b/stage0/stdlib/Lean/Server/FileWorker.c @@ -66,6 +66,7 @@ static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6___closed__ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__25; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initAndRunWorker___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker_mkLspOutputChannel___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_check_canceled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__7___boxed(lean_object*, lean_object*, lean_object*); @@ -197,6 +198,7 @@ static lean_object* l_IO_FS_Stream_readLspNotificationAs___at_Lean_Server_FileWo LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcConnect___rarg(lean_object*, lean_object*); lean_object* l_IO_CancelToken_new(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); lean_object* lean_io_mono_ms_now(lean_object*); lean_object* lean_task_pure(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCancelRequest___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -219,7 +221,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initAndRunWorker_writeErrorDia static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker___hyg_303____closed__2; static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__5; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleStaleDependency(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6___closed__3; static lean_object* l_Lean_Server_FileWorker_parseParams___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__5___boxed(lean_object*); @@ -258,6 +259,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileW LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2470_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleRequest___closed__9; +lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleRpcRelease___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -500,11 +502,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileW LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker___hyg_303_(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_889_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); lean_object* l_Std_Mutex_atomically___at_Std_Channel_recvAllCurrent___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__56; lean_object* lean_string_append(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_FileWorker_runRefreshTask___spec__2(lean_object*); lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1472_(lean_object*); static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_FileWorker_runRefreshTask___spec__2___closed__2; @@ -11565,7 +11565,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_F { lean_object* x_5; lean_inc(x_1); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(x_1); +x_5 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(x_1); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; @@ -15590,7 +15590,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40 lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_165; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_3); x_165 = lean_ctor_get(x_1, 1); lean_inc(x_165); if (lean_obj_tag(x_165) == 0) @@ -15609,7 +15609,7 @@ if (x_167 == 0) { lean_object* x_168; lean_object* x_169; x_168 = lean_ctor_get(x_165, 0); -x_169 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_168); +x_169 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_168); lean_ctor_set(x_165, 0, x_169); x_5 = x_165; x_6 = x_2; @@ -15621,7 +15621,7 @@ lean_object* x_170; lean_object* x_171; lean_object* x_172; x_170 = lean_ctor_get(x_165, 0); lean_inc(x_170); lean_dec(x_165); -x_171 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_170); +x_171 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_170); x_172 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_172, 0, x_171); x_5 = x_172; @@ -25826,7 +25826,7 @@ x_909 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_909, 0, x_907); lean_ctor_set(x_909, 1, x_908); x_910 = l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__24; -x_911 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_910, x_904); +x_911 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_910, x_904); lean_dec(x_904); switch (lean_obj_tag(x_901)) { case 0: @@ -27222,7 +27222,7 @@ x_385 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_385, 0, x_383); lean_ctor_set(x_385, 1, x_384); x_386 = l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__24; -x_387 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_386, x_380); +x_387 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_386, x_380); lean_dec(x_380); switch (lean_obj_tag(x_377)) { case 0: diff --git a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c index 390bd6721b74..fe4458f93206 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c @@ -83,6 +83,8 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePla static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__1; +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(lean_object*); static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__4___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__21___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); @@ -203,7 +205,6 @@ static lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap___closed__1 LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__17___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges_popRanges___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__27___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectInfoBasedSemanticTokens___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__3; static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__10; @@ -254,6 +255,7 @@ lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__4; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__8___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__19(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_findInfoTreeAtPosWithTrailingWhitespace(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__17; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__7(lean_object*, lean_object*, lean_object*); @@ -280,8 +282,10 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__ lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2889_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___rarg(lean_object*, lean_object*); lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t, uint8_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_(lean_object*); static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__4; size_t lean_usize_of_nat(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_pop(lean_object*); static lean_object* l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken___closed__1; @@ -354,7 +358,6 @@ static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorke LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_References_allRefsFor___spec__1(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__6; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1692_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___lambda__1___boxed(lean_object*); @@ -362,7 +365,6 @@ static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_Fil static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__20; static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__3; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); @@ -424,7 +426,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Serve LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1344_(lean_object*); lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____closed__16; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -485,7 +486,6 @@ lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__11(lean_object*); lean_object* l_Lean_Server_locationLinksFromDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__19; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Ord_instDecidableRelLe___rarg(lean_object*, lean_object*, lean_object*); @@ -606,7 +606,6 @@ static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_Fil LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__27___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_noHighlightKinds; static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__1; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__24___lambda__2___closed__2; lean_object* lean_string_length(lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10625_(lean_object*); @@ -706,7 +705,6 @@ lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__37___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__10___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__1___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainGoal(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_handleWaitForDiagnostics___spec__1(lean_object*, lean_object*, lean_object*); @@ -804,6 +802,7 @@ static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_S LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleHover___spec__1___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__3; @@ -836,6 +835,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Serve uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Lean_Server_RequestError_ofIoError(lean_object*); static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____spec__17___lambda__3___closed__1; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12951____closed__28; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__1; static lean_object* l_Lean_Server_FileWorker_filterDuplicateSemanticTokens___closed__1; @@ -23877,7 +23877,7 @@ x_5 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); x_6 = lean_ctor_get(x_2, 0); x_7 = lean_ctor_get(x_2, 1); x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_3, x_6); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_3, x_6); if (x_9 == 0) { uint8_t x_10; @@ -23887,7 +23887,7 @@ return x_10; else { uint8_t x_11; -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_4, x_7); +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_4, x_7); if (x_11 == 0) { uint8_t x_12; @@ -23938,9 +23938,9 @@ x_2 = lean_ctor_get(x_1, 0); x_3 = lean_ctor_get(x_1, 1); x_4 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); x_5 = 0; -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_2); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_2); x_7 = lean_uint64_mix_hash(x_5, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_3); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_180_(x_3); x_9 = lean_uint64_mix_hash(x_7, x_8); x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8939_(x_4); x_11 = lean_uint64_mix_hash(x_9, x_10); @@ -24200,7 +24200,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__L lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24238,7 +24238,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__14; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -24376,7 +24376,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__L lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_2); x_4 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); @@ -24387,7 +24387,7 @@ lean_ctor_set(x_7, 0, x_5); lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_8); x_10 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8962____closed__14; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -24415,7 +24415,7 @@ x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_7); lean_ctor_set(x_20, 1, x_19); x_21 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; -x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); +x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; } @@ -25878,7 +25878,7 @@ x_8 = l_Ord_instDecidableRelLt___rarg(x_7, x_3, x_5); if (x_8 == 0) { uint8_t x_9; -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(x_3, x_5); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_41_(x_3, x_5); lean_dec(x_5); lean_dec(x_3); if (x_9 == 0) @@ -31805,7 +31805,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileW { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -31918,7 +31918,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1155_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); diff --git a/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c b/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c index 3e6b14d0ba5a..574afa69119f 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c +++ b/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c @@ -47,6 +47,7 @@ static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Wi LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389_(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2944_(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instInhabitedMsgToInteractive___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1945____lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__2___closed__2; @@ -165,7 +166,6 @@ static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Wi static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1429____closed__2; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1945____spec__3___closed__2; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2(lean_object*, lean_object*); extern lean_object* l_Lean_Lsp_instFromJsonLocationLink; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__2(lean_object*, lean_object*); @@ -769,7 +769,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -803,7 +803,7 @@ x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_21); lean_ctor_set(x_26, 1, x_25); x_27 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } @@ -1710,7 +1710,7 @@ x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_8); lean_ctor_set(x_17, 1, x_16); x_18 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_17, x_18); +x_19 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_17, x_18); x_20 = l_Lean_Json_mkObj(x_19); return x_20; } @@ -4400,7 +4400,7 @@ x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); lean_ctor_set(x_5, 1, x_4); x_6 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_5, x_6); +x_7 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_5, x_6); x_8 = l_Lean_Json_mkObj(x_7); return x_8; } @@ -4635,7 +4635,7 @@ x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_11); x_13 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); +x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_12, x_13); x_15 = l_Lean_Json_mkObj(x_14); return x_15; } @@ -4669,7 +4669,7 @@ x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_21); lean_ctor_set(x_26, 1, x_25); x_27 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_26, x_27); +x_28 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_26, x_27); x_29 = l_Lean_Json_mkObj(x_28); return x_29; } diff --git a/stage0/stdlib/Lean/Server/GoTo.c b/stage0/stdlib/Lean/Server/GoTo.c index 60fa1e80a30c..131e07e2b8cf 100644 --- a/stage0/stdlib/Lean/Server/GoTo.c +++ b/stage0/stdlib/Lean/Server/GoTo.c @@ -952,7 +952,6 @@ lean_dec(x_7); x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); -lean_inc(x_1); x_11 = l_Lean_isRec___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___spec__1(x_1, x_2, x_3, x_4, x_5, x_9); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); diff --git a/stage0/stdlib/Lean/Server/References.c b/stage0/stdlib/Lean/Server/References.c index d6a4bba9ecd5..0f280b55dbca 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -35,7 +35,6 @@ LEAN_EXPORT lean_object* l_Lean_isRec___at_Lean_Server_RefInfo_toLspRefInfo___sp LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_combineIdents_buildIdMap___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_throwServerError___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_References_findAt(lean_object*, lean_object*, lean_object*, uint8_t); -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_References_allRefs___spec__10(lean_object*, size_t, size_t, lean_object*); static lean_object* l_List_mapTR_loop___at___private_Lean_Server_References_0__Lean_Server_toJsonIlean____x40_Lean_Server_References___hyg_1478____spec__2___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_combineIdents_useConstRepresentatives___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -216,7 +215,6 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_References static lean_object* l_Lean_Server_instFromJsonIlean___closed__1; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_RefInfo_toLspRefInfo___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_References_definitionOf_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_findModuleRefs___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_References_allRefs___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Lsp_ModuleRefs_findAt___spec__1(lean_object*, lean_object*); @@ -333,6 +331,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_combineIdents_buildIdMap(lean_object*, le static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_References_allRefsFor___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_References_findAt___lambda__1___boxed(lean_object*); lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg(lean_object*, lean_object*, lean_object*); +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(lean_object*); lean_object* l_StateT_instMonad___rarg(lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_instRpcEncodableArray___spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_combineIdents_findCanonicalRepresentative___boxed(lean_object*, lean_object*); @@ -376,6 +375,7 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Lsp_ModuleRefs_findRang lean_object* l_Lean_Environment_allImportedModuleNames(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_References_definitionOf_x3f___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_combineIdents_useConstRepresentatives___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_References_removeIlean___spec__5___boxed(lean_object*, lean_object*); lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Lsp_instToJsonModuleRefs___spec__4(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_References_referringTo___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -856,7 +856,6 @@ lean_dec(x_5); x_8 = lean_ctor_get(x_6, 0); lean_inc(x_8); lean_dec(x_6); -lean_inc(x_1); x_9 = l_Lean_isRec___at_Lean_Server_RefInfo_toLspRefInfo___spec__4(x_1, x_2, x_3, x_7); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); @@ -1703,6 +1702,7 @@ lean_object* x_5; x_5 = l_Lean_isRec___at_Lean_Server_RefInfo_toLspRefInfo___spec__4(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); return x_5; } } @@ -12916,7 +12916,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_getModuleContainingDecl_x3f(lean_object* _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); @@ -12933,7 +12933,7 @@ return x_7; else { lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_1, 4); +x_8 = lean_ctor_get(x_1, 5); lean_inc(x_8); lean_dec(x_1); x_9 = lean_ctor_get(x_8, 0); @@ -12992,7 +12992,7 @@ lean_dec(x_1); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); lean_dec(x_7); -x_9 = lean_ctor_get(x_8, 4); +x_9 = lean_ctor_get(x_8, 5); lean_inc(x_9); lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); @@ -13118,7 +13118,7 @@ lean_dec(x_1); x_49 = lean_ctor_get(x_48, 0); lean_inc(x_49); lean_dec(x_48); -x_50 = lean_ctor_get(x_49, 4); +x_50 = lean_ctor_get(x_49, 5); lean_inc(x_50); lean_dec(x_49); x_51 = lean_ctor_get(x_50, 0); @@ -17189,7 +17189,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); x_6 = lean_ctor_get(x_2, 2); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_1); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_1); if (x_7 == 0) { x_2 = x_6; @@ -17231,7 +17231,7 @@ lean_inc(x_13); lean_dec(x_11); x_14 = lean_ctor_get(x_2, 1); x_15 = lean_array_get_size(x_14); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_13); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_13); x_17 = 32; x_18 = lean_uint64_shift_right(x_16, x_17); x_19 = lean_uint64_xor(x_16, x_18); @@ -17680,7 +17680,7 @@ lean_dec(x_1); x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); lean_dec(x_6); -x_8 = lean_ctor_get(x_7, 4); +x_8 = lean_ctor_get(x_7, 5); lean_inc(x_8); lean_dec(x_7); x_9 = lean_ctor_get(x_8, 0); @@ -17880,7 +17880,7 @@ else lean_object* x_4; lean_object* x_5; uint8_t x_6; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_1); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_4, x_1); if (x_6 == 0) { x_2 = x_5; @@ -17912,7 +17912,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8 x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); x_6 = lean_array_get_size(x_1); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_4); +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_4); x_8 = 32; x_9 = lean_uint64_shift_right(x_7, x_8); x_10 = lean_uint64_xor(x_7, x_9); @@ -17943,7 +17943,7 @@ lean_inc(x_23); lean_inc(x_22); lean_dec(x_2); x_25 = lean_array_get_size(x_1); -x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_22); +x_26 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_22); x_27 = 32; x_28 = lean_uint64_shift_right(x_26, x_27); x_29 = lean_uint64_xor(x_26, x_28); @@ -18035,7 +18035,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = lean_ctor_get(x_3, 0); x_7 = lean_ctor_get(x_3, 1); x_8 = lean_ctor_get(x_3, 2); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_6, x_1); +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_6, x_1); if (x_9 == 0) { lean_object* x_10; @@ -18062,7 +18062,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_dec(x_3); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_11, x_1); +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_11, x_1); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; @@ -18127,7 +18127,7 @@ lean_object* x_17; lean_object* x_18; lean_object* x_19; uint64_t x_20; uint64_t x_17 = lean_ctor_get(x_7, 0); x_18 = lean_ctor_get(x_7, 1); x_19 = lean_array_get_size(x_18); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_15); +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_15); x_21 = 32; x_22 = lean_uint64_shift_right(x_20, x_21); x_23 = lean_uint64_xor(x_20, x_22); @@ -18204,7 +18204,7 @@ lean_inc(x_55); lean_inc(x_54); lean_dec(x_7); x_56 = lean_array_get_size(x_55); -x_57 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_15); +x_57 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_15); x_58 = 32; x_59 = lean_uint64_shift_right(x_57, x_58); x_60 = lean_uint64_xor(x_57, x_59); @@ -18514,7 +18514,7 @@ goto _start; else { uint8_t x_18; -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_13, x_15); +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_13, x_15); if (x_18 == 0) { x_2 = x_5; @@ -18676,7 +18676,7 @@ lean_inc(x_28); x_29 = lean_ctor_get(x_26, 1); lean_inc(x_29); lean_dec(x_26); -x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_29); +x_30 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_29); lean_dec(x_29); if (lean_obj_tag(x_28) == 0) { @@ -18854,7 +18854,7 @@ return x_3; else { uint8_t x_21; -x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_16, x_18); +x_21 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_16, x_18); lean_dec(x_18); lean_dec(x_16); if (x_21 == 0) @@ -18939,7 +18939,7 @@ return x_39; else { uint8_t x_40; -x_40 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_34, x_36); +x_40 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_34, x_36); lean_dec(x_36); lean_dec(x_34); if (x_40 == 0) @@ -19010,7 +19010,7 @@ goto _start; else { uint8_t x_19; -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_14, x_16); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_442_(x_14, x_16); if (x_19 == 0) { x_2 = x_6; @@ -19080,7 +19080,7 @@ x_190 = lean_ctor_get(x_10, 1); lean_inc(x_190); x_191 = lean_array_get_size(x_190); x_192 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_158_(x_13); -x_193 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_14); +x_193 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_14); x_194 = 32; x_195 = 16; x_196 = lean_usize_of_nat(x_191); @@ -19213,7 +19213,7 @@ if (lean_is_exclusive(x_10)) { x_23 = lean_array_get_size(x_21); x_24 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_158_(x_13); lean_dec(x_13); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_14); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_14); lean_dec(x_14); x_26 = 32; x_27 = 16; @@ -19366,7 +19366,7 @@ if (lean_is_exclusive(x_10)) { x_80 = lean_array_get_size(x_78); x_81 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_158_(x_13); lean_dec(x_13); -x_82 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_14); +x_82 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_14); lean_dec(x_14); x_83 = 32; x_84 = 16; @@ -19535,7 +19535,7 @@ if (lean_is_exclusive(x_10)) { x_141 = lean_array_get_size(x_139); x_142 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_158_(x_13); lean_dec(x_13); -x_143 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_14); +x_143 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_516_(x_14); lean_dec(x_14); x_144 = 32; x_145 = 16; diff --git a/stage0/stdlib/Lean/Server/Rpc/Deriving.c b/stage0/stdlib/Lean/Server/Rpc/Deriving.c index 4cc8b3337af0..5ca0c8166b09 100644 --- a/stage0/stdlib/Lean/Server/Rpc/Deriving.c +++ b/stage0/stdlib/Lean/Server/Rpc/Deriving.c @@ -97,7 +97,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__7(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__61; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__55; @@ -343,7 +343,7 @@ lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__24; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; LEAN_EXPORT uint8_t l_Lean_Server_RpcEncodable_isOptField(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(size_t, size_t, lean_object*); @@ -1265,7 +1265,8 @@ x_40 = lean_ctor_get(x_37, 1); x_41 = lean_ctor_get(x_39, 0); lean_inc(x_41); lean_dec(x_39); -x_42 = lean_environment_main_module(x_41); +x_42 = l_Lean_Environment_mainModule(x_41); +lean_dec(x_41); x_43 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__3; lean_inc(x_36); x_44 = l_Lean_addMacroScope(x_42, x_43, x_36); @@ -1289,7 +1290,8 @@ x_53 = lean_ctor_get(x_50, 1); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_57 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_35); @@ -1375,7 +1377,8 @@ x_90 = lean_ctor_get(x_87, 1); x_91 = lean_ctor_get(x_89, 0); lean_inc(x_91); lean_dec(x_89); -x_92 = lean_environment_main_module(x_91); +x_92 = l_Lean_Environment_mainModule(x_91); +lean_dec(x_91); x_93 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__43; lean_inc(x_36); lean_inc(x_92); @@ -1435,7 +1438,8 @@ lean_dec(x_87); x_113 = lean_ctor_get(x_111, 0); lean_inc(x_113); lean_dec(x_111); -x_114 = lean_environment_main_module(x_113); +x_114 = l_Lean_Environment_mainModule(x_113); +lean_dec(x_113); x_115 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__43; lean_inc(x_36); lean_inc(x_114); @@ -1497,7 +1501,8 @@ lean_dec(x_50); x_136 = lean_ctor_get(x_134, 0); lean_inc(x_136); lean_dec(x_134); -x_137 = lean_environment_main_module(x_136); +x_137 = l_Lean_Environment_mainModule(x_136); +lean_dec(x_136); x_138 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_139 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_35); @@ -1589,7 +1594,8 @@ if (lean_is_exclusive(x_170)) { x_174 = lean_ctor_get(x_171, 0); lean_inc(x_174); lean_dec(x_171); -x_175 = lean_environment_main_module(x_174); +x_175 = l_Lean_Environment_mainModule(x_174); +lean_dec(x_174); x_176 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__43; lean_inc(x_36); lean_inc(x_175); @@ -1655,7 +1661,8 @@ lean_dec(x_37); x_197 = lean_ctor_get(x_195, 0); lean_inc(x_197); lean_dec(x_195); -x_198 = lean_environment_main_module(x_197); +x_198 = l_Lean_Environment_mainModule(x_197); +lean_dec(x_197); x_199 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__3; lean_inc(x_36); x_200 = l_Lean_addMacroScope(x_198, x_199, x_36); @@ -1685,7 +1692,8 @@ if (lean_is_exclusive(x_206)) { x_210 = lean_ctor_get(x_207, 0); lean_inc(x_210); lean_dec(x_207); -x_211 = lean_environment_main_module(x_210); +x_211 = l_Lean_Environment_mainModule(x_210); +lean_dec(x_210); x_212 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_213 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_35); @@ -1782,7 +1790,8 @@ if (lean_is_exclusive(x_245)) { x_249 = lean_ctor_get(x_246, 0); lean_inc(x_249); lean_dec(x_246); -x_250 = lean_environment_main_module(x_249); +x_250 = l_Lean_Environment_mainModule(x_249); +lean_dec(x_249); x_251 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__43; lean_inc(x_36); lean_inc(x_250); @@ -1857,7 +1866,8 @@ x_277 = lean_ctor_get(x_274, 1); x_278 = lean_ctor_get(x_276, 0); lean_inc(x_278); lean_dec(x_276); -x_279 = lean_environment_main_module(x_278); +x_279 = l_Lean_Environment_mainModule(x_278); +lean_dec(x_278); x_280 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__49; lean_inc(x_273); lean_inc(x_279); @@ -1899,7 +1909,8 @@ x_299 = lean_ctor_get(x_296, 1); x_300 = lean_ctor_get(x_298, 0); lean_inc(x_300); lean_dec(x_298); -x_301 = lean_environment_main_module(x_300); +x_301 = l_Lean_Environment_mainModule(x_300); +lean_dec(x_300); x_302 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_272); x_303 = lean_alloc_ctor(1, 3, 0); @@ -2012,7 +2023,8 @@ x_345 = lean_ctor_get(x_342, 1); x_346 = lean_ctor_get(x_344, 0); lean_inc(x_346); lean_dec(x_344); -x_347 = lean_environment_main_module(x_346); +x_347 = l_Lean_Environment_mainModule(x_346); +lean_dec(x_346); lean_inc(x_273); lean_inc(x_347); x_348 = l_Lean_addMacroScope(x_347, x_310, x_273); @@ -2086,7 +2098,8 @@ lean_dec(x_342); x_372 = lean_ctor_get(x_370, 0); lean_inc(x_372); lean_dec(x_370); -x_373 = lean_environment_main_module(x_372); +x_373 = l_Lean_Environment_mainModule(x_372); +lean_dec(x_372); lean_inc(x_273); lean_inc(x_373); x_374 = l_Lean_addMacroScope(x_373, x_310, x_273); @@ -2162,7 +2175,8 @@ lean_dec(x_296); x_399 = lean_ctor_get(x_397, 0); lean_inc(x_399); lean_dec(x_397); -x_400 = lean_environment_main_module(x_399); +x_400 = l_Lean_Environment_mainModule(x_399); +lean_dec(x_399); x_401 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_272); x_402 = lean_alloc_ctor(1, 3, 0); @@ -2281,7 +2295,8 @@ if (lean_is_exclusive(x_442)) { x_446 = lean_ctor_get(x_443, 0); lean_inc(x_446); lean_dec(x_443); -x_447 = lean_environment_main_module(x_446); +x_447 = l_Lean_Environment_mainModule(x_446); +lean_dec(x_446); lean_inc(x_273); lean_inc(x_447); x_448 = l_Lean_addMacroScope(x_447, x_410, x_273); @@ -2361,7 +2376,8 @@ lean_dec(x_274); x_473 = lean_ctor_get(x_471, 0); lean_inc(x_473); lean_dec(x_471); -x_474 = lean_environment_main_module(x_473); +x_474 = l_Lean_Environment_mainModule(x_473); +lean_dec(x_473); x_475 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__49; lean_inc(x_273); lean_inc(x_474); @@ -2409,7 +2425,8 @@ if (lean_is_exclusive(x_491)) { x_495 = lean_ctor_get(x_492, 0); lean_inc(x_495); lean_dec(x_492); -x_496 = lean_environment_main_module(x_495); +x_496 = l_Lean_Environment_mainModule(x_495); +lean_dec(x_495); x_497 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_272); x_498 = lean_alloc_ctor(1, 3, 0); @@ -2533,7 +2550,8 @@ if (lean_is_exclusive(x_539)) { x_543 = lean_ctor_get(x_540, 0); lean_inc(x_543); lean_dec(x_540); -x_544 = lean_environment_main_module(x_543); +x_544 = l_Lean_Environment_mainModule(x_543); +lean_dec(x_543); lean_inc(x_273); lean_inc(x_544); x_545 = l_Lean_addMacroScope(x_544, x_507, x_273); @@ -4228,7 +4246,8 @@ x_44 = lean_ctor_get(x_42, 0); x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); lean_dec(x_44); -x_46 = lean_environment_main_module(x_45); +x_46 = l_Lean_Environment_mainModule(x_45); +lean_dec(x_45); x_47 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_48 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_40); @@ -4754,7 +4773,8 @@ lean_dec(x_42); x_268 = lean_ctor_get(x_266, 0); lean_inc(x_268); lean_dec(x_266); -x_269 = lean_environment_main_module(x_268); +x_269 = l_Lean_Environment_mainModule(x_268); +lean_dec(x_268); x_270 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_271 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_40); @@ -5347,7 +5367,8 @@ if (lean_is_exclusive(x_504)) { x_508 = lean_ctor_get(x_505, 0); lean_inc(x_508); lean_dec(x_505); -x_509 = lean_environment_main_module(x_508); +x_509 = l_Lean_Environment_mainModule(x_508); +lean_dec(x_508); x_510 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_511 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_502); @@ -5955,7 +5976,8 @@ if (lean_is_exclusive(x_747)) { x_751 = lean_ctor_get(x_748, 0); lean_inc(x_751); lean_dec(x_748); -x_752 = lean_environment_main_module(x_751); +x_752 = l_Lean_Environment_mainModule(x_751); +lean_dec(x_751); x_753 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_754 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_745); @@ -6578,7 +6600,8 @@ if (lean_is_exclusive(x_993)) { x_997 = lean_ctor_get(x_994, 0); lean_inc(x_997); lean_dec(x_994); -x_998 = lean_environment_main_module(x_997); +x_998 = l_Lean_Environment_mainModule(x_997); +lean_dec(x_997); x_999 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_1000 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_991); @@ -7216,7 +7239,8 @@ if (lean_is_exclusive(x_1242)) { x_1246 = lean_ctor_get(x_1243, 0); lean_inc(x_1246); lean_dec(x_1243); -x_1247 = lean_environment_main_module(x_1246); +x_1247 = l_Lean_Environment_mainModule(x_1246); +lean_dec(x_1246); x_1248 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_1249 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_1240); @@ -8588,7 +8612,8 @@ x_27 = lean_ctor_get(x_24, 1); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__59; lean_inc(x_22); lean_ctor_set_tag(x_24, 2); @@ -8647,7 +8672,8 @@ lean_dec(x_24); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__59; lean_inc(x_22); x_57 = lean_alloc_ctor(2, 2, 0); @@ -8963,7 +8989,8 @@ x_26 = lean_ctor_get(x_23, 1); x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); -x_28 = lean_environment_main_module(x_27); +x_28 = l_Lean_Environment_mainModule(x_27); +lean_dec(x_27); x_29 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; lean_inc(x_21); lean_ctor_set_tag(x_23, 2); @@ -9010,7 +9037,8 @@ lean_dec(x_23); x_47 = lean_ctor_get(x_45, 0); lean_inc(x_47); lean_dec(x_45); -x_48 = lean_environment_main_module(x_47); +x_48 = l_Lean_Environment_mainModule(x_47); +lean_dec(x_47); x_49 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; lean_inc(x_21); x_50 = lean_alloc_ctor(2, 2, 0); @@ -9090,7 +9118,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_29); +x_30 = l_Lean_Environment_mainModule(x_29); +lean_dec(x_29); x_31 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; lean_inc(x_23); lean_ctor_set_tag(x_25, 2); @@ -9141,7 +9170,8 @@ lean_dec(x_25); x_49 = lean_ctor_get(x_47, 0); lean_inc(x_49); lean_dec(x_47); -x_50 = lean_environment_main_module(x_49); +x_50 = l_Lean_Environment_mainModule(x_49); +lean_dec(x_49); x_51 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; lean_inc(x_23); x_52 = lean_alloc_ctor(2, 2, 0); @@ -9411,7 +9441,8 @@ x_27 = lean_ctor_get(x_24, 1); x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_environment_main_module(x_28); +x_29 = l_Lean_Environment_mainModule(x_28); +lean_dec(x_28); x_30 = lean_box(0); x_31 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; @@ -9506,7 +9537,8 @@ x_73 = lean_ctor_get(x_70, 1); x_74 = lean_ctor_get(x_72, 0); lean_inc(x_74); lean_dec(x_72); -x_75 = lean_environment_main_module(x_74); +x_75 = l_Lean_Environment_mainModule(x_74); +lean_dec(x_74); x_76 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; lean_inc(x_22); lean_ctor_set_tag(x_70, 2); @@ -9741,7 +9773,8 @@ lean_dec(x_70); x_157 = lean_ctor_get(x_155, 0); lean_inc(x_157); lean_dec(x_155); -x_158 = lean_environment_main_module(x_157); +x_158 = l_Lean_Environment_mainModule(x_157); +lean_dec(x_157); x_159 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; lean_inc(x_22); x_160 = lean_alloc_ctor(2, 2, 0); @@ -9924,7 +9957,8 @@ if (lean_is_exclusive(x_216)) { x_220 = lean_ctor_get(x_217, 0); lean_inc(x_220); lean_dec(x_217); -x_221 = lean_environment_main_module(x_220); +x_221 = l_Lean_Environment_mainModule(x_220); +lean_dec(x_220); x_222 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; lean_inc(x_22); if (lean_is_scalar(x_219)) { @@ -10132,7 +10166,8 @@ lean_dec(x_24); x_284 = lean_ctor_get(x_282, 0); lean_inc(x_284); lean_dec(x_282); -x_285 = lean_environment_main_module(x_284); +x_285 = l_Lean_Environment_mainModule(x_284); +lean_dec(x_284); x_286 = lean_box(0); x_287 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; x_288 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; @@ -10239,7 +10274,8 @@ if (lean_is_exclusive(x_327)) { x_331 = lean_ctor_get(x_328, 0); lean_inc(x_331); lean_dec(x_328); -x_332 = lean_environment_main_module(x_331); +x_332 = l_Lean_Environment_mainModule(x_331); +lean_dec(x_331); x_333 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; lean_inc(x_22); if (lean_is_scalar(x_330)) { @@ -11298,7 +11334,8 @@ x_57 = lean_ctor_get(x_55, 0); x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); lean_dec(x_57); -x_59 = lean_environment_main_module(x_58); +x_59 = l_Lean_Environment_mainModule(x_58); +lean_dec(x_58); lean_inc(x_33); x_60 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_60, 0, x_33); @@ -11850,7 +11887,8 @@ lean_dec(x_55); x_295 = lean_ctor_get(x_293, 0); lean_inc(x_295); lean_dec(x_293); -x_296 = lean_environment_main_module(x_295); +x_296 = l_Lean_Environment_mainModule(x_295); +lean_dec(x_295); lean_inc(x_33); x_297 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_297, 0, x_33); @@ -12451,7 +12489,8 @@ if (lean_is_exclusive(x_550)) { x_554 = lean_ctor_get(x_551, 0); lean_inc(x_554); lean_dec(x_551); -x_555 = lean_environment_main_module(x_554); +x_555 = l_Lean_Environment_mainModule(x_554); +lean_dec(x_554); lean_inc(x_33); x_556 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_556, 0, x_33); @@ -13123,7 +13162,8 @@ if (lean_is_exclusive(x_824)) { x_828 = lean_ctor_get(x_825, 0); lean_inc(x_828); lean_dec(x_825); -x_829 = lean_environment_main_module(x_828); +x_829 = l_Lean_Environment_mainModule(x_828); +lean_dec(x_828); lean_inc(x_802); x_830 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_830, 0, x_802); @@ -13810,7 +13850,8 @@ if (lean_is_exclusive(x_1103)) { x_1107 = lean_ctor_get(x_1104, 0); lean_inc(x_1107); lean_dec(x_1104); -x_1108 = lean_environment_main_module(x_1107); +x_1108 = l_Lean_Environment_mainModule(x_1107); +lean_dec(x_1107); lean_inc(x_1081); x_1109 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1109, 0, x_1081); @@ -14961,8 +15002,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_environment_find(x_9, x_1); +x_10 = l_Lean_Environment_find_x3f(x_9, x_1); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -15003,8 +15043,7 @@ lean_dec(x_5); x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_1); -x_22 = lean_environment_find(x_21, x_1); +x_22 = l_Lean_Environment_find_x3f(x_21, x_1); if (lean_obj_tag(x_22) == 0) { uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -15342,7 +15381,8 @@ x_26 = lean_ctor_get(x_23, 1); x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); -x_28 = lean_environment_main_module(x_27); +x_28 = l_Lean_Environment_mainModule(x_27); +lean_dec(x_27); x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__5___closed__3; lean_inc(x_21); lean_ctor_set_tag(x_23, 2); @@ -15399,7 +15439,8 @@ lean_dec(x_23); x_53 = lean_ctor_get(x_51, 0); lean_inc(x_53); lean_dec(x_51); -x_54 = lean_environment_main_module(x_53); +x_54 = l_Lean_Environment_mainModule(x_53); +lean_dec(x_53); x_55 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__5___closed__3; lean_inc(x_21); x_56 = lean_alloc_ctor(2, 2, 0); diff --git a/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c b/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c index e9a593c531ea..3263d8d1ef07 100644 --- a/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c @@ -201,7 +201,7 @@ static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling_ static lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___closed__2; static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1385____closed__6; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); 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*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__2___closed__3; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__2___closed__6; @@ -3335,7 +3335,8 @@ x_18 = lean_ctor_get(x_15, 1); x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_environment_main_module(x_19); +x_20 = l_Lean_Environment_mainModule(x_19); +lean_dec(x_19); x_21 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__7; x_22 = l_Lean_addMacroScope(x_20, x_21, x_14); x_23 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__8; @@ -3623,7 +3624,8 @@ lean_dec(x_15); x_102 = lean_ctor_get(x_100, 0); lean_inc(x_102); lean_dec(x_100); -x_103 = lean_environment_main_module(x_102); +x_103 = l_Lean_Environment_mainModule(x_102); +lean_dec(x_102); x_104 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__7; x_105 = l_Lean_addMacroScope(x_103, x_104, x_14); x_106 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__8; diff --git a/stage0/stdlib/Lean/Server/Watchdog.c b/stage0/stdlib/Lean/Server/Watchdog.c index df9fb0e5efa2..0b97fb0addb0 100644 --- a/stage0/stdlib/Lean/Server/Watchdog.c +++ b/stage0/stdlib/Lean/Server/Watchdog.c @@ -60,6 +60,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_queuedMsgs___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_ImportData_update___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handlePrepareCallHierarchy___spec__3___closed__1; static lean_object* l_Lean_Server_Watchdog_FileWorker_queuedMsgs___closed__1; @@ -71,6 +72,7 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handlePrep static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__40; static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__16; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); static size_t l_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleNotification___spec__14___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls_collapseSameOutgoingCalls___boxed(lean_object*); @@ -88,7 +90,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleReference(lean_object*, le static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__13; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_handleDidSave___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleRename(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls___spec__5___lambda__1___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_309_(lean_object*); @@ -167,7 +168,6 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Server_Watchdog_handleCanc lean_object* lean_io_getenv(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__24; uint8_t lean_float_decLt(double, double); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__15; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__2___closed__4; lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_References_allRefsFor___spec__5(lean_object*, lean_object*); @@ -266,6 +266,7 @@ uint8_t lean_string_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_ImportData_update___spec__6(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__12(size_t, size_t, lean_object*); +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_waitForProc(lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_FileWorker_erasePendingRequest___spec__2(lean_object*, lean_object*); @@ -294,7 +295,6 @@ static lean_object* l_Lean_Server_Watchdog_handleNotification___closed__2; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4173____closed__15; lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Lsp_Ipc_shutdown___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleNotification___spec__9(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_Watchdog_handleRename___spec__11(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__2; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4173____closed__2; @@ -328,7 +328,6 @@ static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__1; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4173____closed__10; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux___closed__10; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_eraseFileWorker(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_eraseFileWorker___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_CrashOrigin_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__11(lean_object*, lean_object*); @@ -345,6 +344,7 @@ lean_object* lean_io_wait_any(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__8(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_Watchdog_ImportData_update___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls_collapseSameOutgoingCalls___spec__5(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__26; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__8(lean_object*, size_t, size_t, lean_object*); @@ -383,6 +383,7 @@ lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTex static lean_object* l_Lean_Server_Watchdog_workerCfg___closed__1; lean_object* l_Lean_FileMap_ofString(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Watchdog_handleRename___spec__13___boxed(lean_object*, lean_object*, lean_object*); static uint8_t l_Lean_Server_Watchdog_tryDischargeQueuedMessages___closed__3; lean_object* lean_array_to_list(lean_object*); @@ -499,7 +500,6 @@ LEAN_EXPORT lean_object* l_IO_FS_Stream_readLspNotificationAs___at_Lean_Server_W static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__27; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux___closed__14; lean_object* l_Lean_Server_References_findRange_x3f(lean_object*, lean_object*, lean_object*, uint8_t); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(lean_object*); static lean_object* l_Lean_Server_Watchdog_mainLoop___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleRequest___spec__14___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_fold___at_Lean_RBTree_diff___spec__1___rarg(lean_object*, lean_object*, lean_object*); @@ -508,7 +508,6 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_ImportData_ LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__4(lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__12; static lean_object* l_Lean_Server_Watchdog_handleNotification___closed__3; -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_findWorkerPath___closed__2; static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleRequest___spec__10(size_t, size_t, lean_object*); @@ -789,12 +788,12 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_startLoadingReferences(lean_obje static lean_object* l_Lean_Server_Watchdog_tryDischargeQueuedMessages___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_killProcAndWait___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_handleDidSave___spec__3___at_Lean_Server_Watchdog_handleDidSave___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(lean_object*, lean_object*); static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__4___closed__1; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__7___closed__1; static lean_object* l_Lean_Server_Watchdog_handleRename___closed__2; static lean_object* l_Lean_Server_Watchdog_handleDidOpen___closed__1; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4173____closed__7; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__19; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_Watchdog_forwardRequestToWorker___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__12___closed__1; @@ -804,7 +803,6 @@ lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistrationPara lean_object* lean_string_append(lean_object*, lean_object*); lean_object* l_Lean_Server_References_addIlean(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__55; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_Watchdog_handleRequest___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleRequest___spec__10___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_callHierarchyItemOf_x3f(lean_object*, lean_object*, lean_object*, lean_object*); @@ -821,6 +819,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_Watchdog_ImportData_ extern lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall; lean_object* lean_int_neg(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Watchdog_findDefinitions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__64; @@ -870,6 +869,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Server_Watchdog_ImportDat LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handlePrepareCallHierarchy___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SearchPath_findAllWithExt(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Watchdog_handleRename___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -25937,7 +25937,7 @@ x_9 = lean_ctor_get(x_1, 0); x_10 = lean_ctor_get(x_1, 1); x_11 = lean_ctor_get(x_1, 2); x_12 = lean_ctor_get(x_1, 3); -x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_2, x_10); +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_2, x_10); switch (x_13) { case 0: { @@ -25982,7 +25982,7 @@ lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); lean_dec(x_1); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_2, x_20); +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_2, x_20); switch (x_23) { case 0: { @@ -26038,7 +26038,7 @@ x_33 = lean_ctor_get(x_1, 0); x_34 = lean_ctor_get(x_1, 1); x_35 = lean_ctor_get(x_1, 2); x_36 = lean_ctor_get(x_1, 3); -x_37 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_2, x_34); +x_37 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_2, x_34); switch (x_37) { case 0: { @@ -27432,7 +27432,7 @@ lean_inc(x_344); lean_inc(x_343); lean_inc(x_342); lean_dec(x_1); -x_346 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300_(x_2, x_343); +x_346 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_115_(x_2, x_343); switch (x_346) { case 0: { @@ -34377,7 +34377,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_ _start: { lean_object* x_2; uint8_t x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(x_1); +x_2 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86_(x_1); x_3 = !lean_is_exclusive(x_2); if (x_3 == 0) { @@ -37875,7 +37875,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_2, 1); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(x_5); +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(x_5); lean_ctor_set_tag(x_2, 2); lean_ctor_set(x_2, 1, x_6); x_7 = l_IO_FS_Stream_writeLspMessage(x_1, x_2, x_3); @@ -37889,7 +37889,7 @@ x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_inc(x_8); lean_dec(x_2); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4408_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4223_(x_9); x_11 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_11, 0, x_8); lean_ctor_set(x_11, 1, x_10); @@ -37984,7 +37984,7 @@ lean_dec(x_14); x_15 = lean_ctor_get(x_4, 0); lean_inc(x_15); lean_dec(x_4); -x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_15); +x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_15); lean_ctor_set_tag(x_2, 2); lean_ctor_set(x_2, 1, x_16); x_17 = l_IO_FS_Stream_writeLspMessage(x_1, x_2, x_3); @@ -37999,7 +37999,7 @@ lean_dec(x_2); x_19 = lean_ctor_get(x_4, 0); lean_inc(x_19); lean_dec(x_4); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_19); +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_19); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_20); @@ -38479,7 +38479,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(x_5); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_895_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -38534,7 +38534,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5326_(x_1); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5141_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -40756,7 +40756,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(x_1); +x_4 = l___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_124_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -45012,7 +45012,7 @@ x_371 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_371, 0, x_369); lean_ctor_set(x_371, 1, x_370); x_372 = l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__14; -x_373 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_372, x_366); +x_373 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_372, x_366); lean_dec(x_366); switch (lean_obj_tag(x_363)) { case 0: @@ -52284,7 +52284,7 @@ x_909 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_909, 0, x_907); lean_ctor_set(x_909, 1, x_908); x_910 = l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__14; -x_911 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_910, x_904); +x_911 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_910, x_904); lean_dec(x_904); switch (lean_obj_tag(x_901)) { case 0: diff --git a/stage0/stdlib/Lean/Structure.c b/stage0/stdlib/Lean/Structure.c index 20960265189a..de54c7fbbb90 100644 --- a/stage0/stdlib/Lean/Structure.c +++ b/stage0/stdlib/Lean/Structure.c @@ -42,6 +42,7 @@ lean_object* l___private_Lean_Expr_0__Lean_reprBinderInfo____x40_Lean_Expr___hyg static lean_object* l___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f___closed__1; static lean_object* l_Lean_initFn____x40_Lean_Structure___hyg_417____closed__3; LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_StructureInfo_getProjFn_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getStructureLikeCtor_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Structure___hyg_417____lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Structure___hyg_417____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -84,7 +85,7 @@ LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_computeStructureResolu static lean_object* l_Lean_initFn____x40_Lean_Structure___hyg_417____closed__7; lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_initFn____x40_Lean_Structure___hyg_417____spec__6(lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_computeStructureResolutionOrder___spec__7(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_setStructureParents___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__2___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -219,6 +220,7 @@ static lean_object* l___private_Lean_Structure_0__Lean_reprStructureFieldInfo___ LEAN_EXPORT lean_object* l_Lean_instInhabitedStructureDescr; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Structure_0__Lean_setStructureResolutionOrder___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getStructureLikeNumFields___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_computeStructureResolutionOrder___spec__14___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_toArray___at_Lean_initFn____x40_Lean_Structure___hyg_417____spec__5___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___rarg___lambda__1___boxed(lean_object*, lean_object*); @@ -330,6 +332,7 @@ static lean_object* l___private_Lean_Structure_0__Lean_reprStructureFieldInfo___ lean_object* l_Array_indexOfAux___at_Lean_MetavarContext_setMVarUserName___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getStructureCtor___closed__3; lean_object* l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getStructureCtor___boxed(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isStructure___boxed(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); @@ -3116,7 +3119,7 @@ LEAN_EXPORT lean_object* l_Lean_getStructureCtor(lean_object* x_1, lean_object* { lean_object* x_3; lean_inc(x_1); -x_3 = lean_environment_find(x_1, x_2); +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_3) == 0) { lean_object* x_4; lean_object* x_5; @@ -3159,7 +3162,8 @@ lean_object* x_12; lean_object* x_13; x_12 = lean_ctor_get(x_8, 0); lean_inc(x_12); lean_dec(x_8); -x_13 = lean_environment_find(x_1, x_12); +x_13 = l_Lean_Environment_find_x3f(x_1, x_12); +lean_dec(x_12); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; lean_object* x_15; @@ -3215,6 +3219,15 @@ return x_23; } } } +LEAN_EXPORT lean_object* l_Lean_getStructureCtor___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_getStructureCtor(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_getStructureFields(lean_object* x_1, lean_object* x_2) { _start: { @@ -4517,7 +4530,7 @@ LEAN_EXPORT uint8_t l_Lean_isStructureLike(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_environment_find(x_1, x_2); +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4606,6 +4619,7 @@ LEAN_EXPORT lean_object* l_Lean_isStructureLike___boxed(lean_object* x_1, lean_o { uint8_t x_3; lean_object* x_4; x_3 = l_Lean_isStructureLike(x_1, x_2); +lean_dec(x_2); x_4 = lean_box(x_3); return x_4; } @@ -4645,7 +4659,7 @@ LEAN_EXPORT lean_object* l_Lean_getStructureLikeCtor_x3f(lean_object* x_1, lean_ { lean_object* x_3; lean_inc(x_1); -x_3 = lean_environment_find(x_1, x_2); +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_3) == 0) { lean_object* x_4; @@ -4704,7 +4718,8 @@ lean_object* x_15; lean_object* x_16; x_15 = lean_ctor_get(x_8, 0); lean_inc(x_15); lean_dec(x_8); -x_16 = lean_environment_find(x_1, x_15); +x_16 = l_Lean_Environment_find_x3f(x_1, x_15); +lean_dec(x_15); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; @@ -4798,12 +4813,21 @@ return x_31; } } } +LEAN_EXPORT lean_object* l_Lean_getStructureLikeCtor_x3f___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_getStructureLikeCtor_x3f(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_getStructureLikeNumFields(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_inc(x_1); -x_3 = lean_environment_find(x_1, x_2); +x_3 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_3) == 0) { lean_object* x_4; @@ -4862,7 +4886,8 @@ lean_object* x_15; lean_object* x_16; x_15 = lean_ctor_get(x_8, 0); lean_inc(x_15); lean_dec(x_8); -x_16 = lean_environment_find(x_1, x_15); +x_16 = l_Lean_Environment_find_x3f(x_1, x_15); +lean_dec(x_15); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; @@ -4927,6 +4952,15 @@ return x_24; } } } +LEAN_EXPORT lean_object* l_Lean_getStructureLikeNumFields___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_getStructureLikeNumFields(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} static lean_object* _init_l_Lean_instInhabitedStructureResolutionState() { _start: { diff --git a/stage0/stdlib/Lean/Util/CollectAxioms.c b/stage0/stdlib/Lean/Util/CollectAxioms.c index cedbbc72b4d4..4e80b96fc3af 100644 --- a/stage0/stdlib/Lean/Util/CollectAxioms.c +++ b/stage0/stdlib/Lean/Util/CollectAxioms.c @@ -16,7 +16,7 @@ extern "C" { LEAN_EXPORT lean_object* l_Lean_CollectAxioms_collect(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_CollectAxioms_collect___spec__2(lean_object*, lean_object*, lean_object*); @@ -129,9 +129,8 @@ x_11 = l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(x_4, x_1, x_10); lean_inc(x_5); lean_inc(x_11); lean_ctor_set(x_3, 0, x_11); -lean_inc(x_1); lean_inc(x_2); -x_12 = lean_environment_find(x_2, x_1); +x_12 = l_Lean_Environment_find_x3f(x_2, x_1); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; @@ -615,9 +614,8 @@ lean_inc(x_123); x_124 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_124, 0, x_123); lean_ctor_set(x_124, 1, x_5); -lean_inc(x_1); lean_inc(x_2); -x_125 = lean_environment_find(x_2, x_1); +x_125 = l_Lean_Environment_find_x3f(x_2, x_1); if (lean_obj_tag(x_125) == 0) { lean_object* x_126; diff --git a/stage0/stdlib/Lean/Util/FoldConsts.c b/stage0/stdlib/Lean/Util/FoldConsts.c index 34bc2f30439d..7effcf100974 100644 --- a/stage0/stdlib/Lean/Util/FoldConsts.c +++ b/stage0/stdlib/Lean/Util/FoldConsts.c @@ -31,7 +31,7 @@ LEAN_EXPORT lean_object* l_Lean_ConstantInfo_getUsedConstantsAsSet(lean_object*) lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_FoldConstsImpl_fold(lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_getUsedConstants___closed__2; @@ -1097,7 +1097,7 @@ LEAN_EXPORT uint32_t l_Lean_getMaxHeight___lambda__1(lean_object* x_1, lean_obje _start: { lean_object* x_4; -x_4 = lean_environment_find(x_1, x_2); +x_4 = l_Lean_Environment_find_x3f(x_1, x_2); if (lean_obj_tag(x_4) == 0) { return x_3; @@ -1177,6 +1177,7 @@ uint32_t x_4; uint32_t x_5; lean_object* x_6; x_4 = lean_unbox_uint32(x_3); lean_dec(x_3); x_5 = l_Lean_getMaxHeight___lambda__1(x_1, x_2, x_4); +lean_dec(x_2); x_6 = lean_box_uint32(x_5); return x_6; } diff --git a/stage0/stdlib/Lean/Util/ShareCommon.c b/stage0/stdlib/Lean/Util/ShareCommon.c index aaf582ebde5e..273cdd620545 100644 --- a/stage0/stdlib/Lean/Util/ShareCommon.c +++ b/stage0/stdlib/Lean/Util/ShareCommon.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Util.ShareCommon -// Imports: Init.ShareCommon Std.Data.HashSet Std.Data.HashMap Lean.Data.PersistentHashMap Lean.Data.PersistentHashSet +// Imports: Init.ShareCommon Std.Data.HashSet.Basic Std.Data.HashMap.Basic Lean.Data.PersistentHashMap Lean.Data.PersistentHashSet #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -1100,8 +1100,8 @@ return x_2; } } lean_object* initialize_Init_ShareCommon(uint8_t builtin, lean_object*); -lean_object* initialize_Std_Data_HashSet(uint8_t builtin, lean_object*); -lean_object* initialize_Std_Data_HashMap(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Data_HashSet_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Data_HashMap_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_PersistentHashMap(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_PersistentHashSet(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -1112,10 +1112,10 @@ _G_initialized = true; res = initialize_Init_ShareCommon(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Std_Data_HashSet(builtin, lean_io_mk_world()); +res = initialize_Std_Data_HashSet_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Std_Data_HashMap(builtin, lean_io_mk_world()); +res = initialize_Std_Data_HashMap_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); res = initialize_Lean_Data_PersistentHashMap(builtin, lean_io_mk_world()); diff --git a/stage0/stdlib/Lean/Util/Trace.c b/stage0/stdlib/Lean/Util/Trace.c index fbda915c6df9..a2f32b862f49 100644 --- a/stage0/stdlib/Lean/Util/Trace.c +++ b/stage0/stdlib/Lean/Util/Trace.c @@ -22,6 +22,7 @@ LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptReaderT___rarg(lean_object* LEAN_EXPORT lean_object* l_Lean_addTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateRefT_x27_instMonadExceptOf___rarg(lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___closed__2; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_addTraceAsMessages___spec__5___at_Lean_addTraceAsMessages___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1644_(lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); @@ -29,6 +30,7 @@ static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__10; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_75_(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__9; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instMonadAlwaysExceptEIO___closed__1; @@ -38,22 +40,23 @@ static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__22; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__23; LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; static lean_object* l_Lean_resetTraceState___rarg___lambda__1___closed__2; lean_object* l_IO_getNumHeartbeats___boxed(lean_object*); lean_object* l_instBEqOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__34; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__6(lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__32; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_75____closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__35; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode_x27___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*); LEAN_EXPORT lean_object* l_Lean_exceptOptionEmoji(lean_object*, lean_object*); static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__4; @@ -63,15 +66,14 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__9(lean_obje static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateT_instMonadExceptOf___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18; lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____closed__1; static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__10(lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__27; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33; static lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec__15___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_instExceptToEmojiOption(lean_object*, lean_object*); @@ -87,13 +89,13 @@ uint64_t lean_uint64_of_nat(lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__43; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__16; lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7; size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31; +static lean_object* l_Lean_expandTraceMacro___closed__15; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_modifyTraces___rarg(lean_object*, lean_object*); @@ -102,12 +104,11 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateRefT_x27(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44; lean_object* l_Lean_Syntax_getId(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__6; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__1; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__2; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__51; lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__2(lean_object*, lean_object*); @@ -127,12 +128,10 @@ static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____closed__3 static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__19; LEAN_EXPORT lean_object* l_Lean_getTraces(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_printTraces___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1460_(lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptMonadCacheT(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -145,15 +144,12 @@ lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__9; static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__2; lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__9; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__6___boxed(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__5___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Lean_traceM(lean_object*); static lean_object* l_Lean_crossEmoji___closed__1; static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__1; @@ -166,7 +162,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__4___boxed(l LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTraceElem___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__15; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__1(lean_object*); @@ -178,8 +173,8 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces__ size_t lean_usize_of_nat(lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); static size_t l_Lean_instInhabitedTraceState___closed__4; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5; LEAN_EXPORT lean_object* l_Lean_trace_profiler_useHeartbeats; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__46; LEAN_EXPORT lean_object* l_Lean_addRawTrace(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -192,7 +187,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessag lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_40____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg(lean_object*); @@ -202,14 +196,13 @@ lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instMonadExceptOfEIO(lean_object*); static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__5; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__3; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368____closed__5; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3___boxed(lean_object**); static lean_object* l_Lean_instInhabitedTraceElem___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_addTraceAsMessages___spec__8(lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__16; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__3___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState(lean_object*); static lean_object* l_Lean_instInhabitedTraceState___closed__2; @@ -218,6 +211,7 @@ static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instExceptToEmojiOption___closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__10; LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces(lean_object*); @@ -225,7 +219,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode(lean_o static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1460____closed__5; LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3; static lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__2; static lean_object* l_Lean_bombEmoji___closed__1; size_t lean_usize_of_nat(lean_object*); @@ -240,8 +233,8 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__13(lean_object*, LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_addTraceAsMessages___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__11___boxed(lean_object**); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__2; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__36; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___rarg___lambda__6___closed__2; @@ -252,7 +245,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__10___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; static lean_object* l_Lean_resetTraceState___rarg___closed__1; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__20; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -266,6 +258,7 @@ LEAN_EXPORT lean_object* l_Lean_exceptEmoji___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_NameSSet_insert___spec__7(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__29; LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); @@ -279,12 +272,12 @@ lean_object* l_instDecidableEqPos___boxed(lean_object*, lean_object*); lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__2; LEAN_EXPORT lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_75____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1644____closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1460____closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_withTraceNode___rarg___lambda__6___closed__1; @@ -293,16 +286,16 @@ LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1___boxed(lean_ lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__12; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14; static lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_printTraces(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10; +static lean_object* l_Lean_expandTraceMacro___closed__12; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6___rarg___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276_(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__42; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__10(lean_object*); lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__2(lean_object*); @@ -318,43 +311,39 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getTraces___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateT(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__8; uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__4___boxed(lean_object**); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__26; LEAN_EXPORT uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption_go(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__12; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__45; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50; lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__15; LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_addTraceAsMessages___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__31; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___at_Lean_addTraceAsMessages___spec__17___boxed(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____closed__2; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Lean_instInhabitedTraceElem; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__14___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24; +static lean_object* l_Lean_expandTraceMacro___closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_print___at_IO_println___spec__1(lean_object*, lean_object*); @@ -368,8 +357,6 @@ uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_IO_println___at_Lean_printTraces___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_addTraceAsMessages___spec__9(lean_object*); @@ -378,24 +365,26 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___sp LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace_profiler_threshold; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; static lean_object* l_Lean_instExceptToEmojiBool___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__2(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__4; +static lean_object* l_Lean_expandTraceMacro___closed__6; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__38; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__8; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__44; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__11; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____closed__4; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__48; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Lean_traceM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25; extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1; LEAN_EXPORT lean_object* l_Lean_trace(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__3; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6; static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__2; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__18; lean_object* lean_register_option(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptMonadCacheT___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__10___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); @@ -409,7 +398,6 @@ static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__12___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__5; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__1; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); @@ -418,12 +406,14 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___sp static lean_object* l_Lean_instInhabitedTraceState___closed__6; static lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___closed__1; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__21; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerTraceClass___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__14; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__47; LEAN_EXPORT lean_object* l_Lean_instMonadTraceOfMonadLift___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__18; LEAN_EXPORT lean_object* l_Lean_trace___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__2; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__24; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368____closed__4; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__4; @@ -438,7 +428,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_addTraceAsMessag LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__17; LEAN_EXPORT lean_object* l_Lean_getTraces___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__2___boxed(lean_object*); @@ -449,12 +438,14 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec_ static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__10; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__27; +static lean_object* l_Lean_expandTraceMacro___closed__16; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptEIO(lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__14; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__7___boxed(lean_object**); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__1(lean_object*, lean_object*, lean_object*); @@ -462,66 +453,73 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessag static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__6; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__41; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__14___rarg___closed__3; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_addTraceAsMessages___spec__7(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__22; static lean_object* l_Lean_addTraceAsMessages___rarg___lambda__5___closed__1; +static lean_object* l_Lean_expandTraceMacro___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateT___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNodeBefore___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__19; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__18; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__14; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__14___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014_; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__7; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__24; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__6___boxed(lean_object**); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__20; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__3___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4___rarg___lambda__3(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__21; LEAN_EXPORT lean_object* l_Lean_exceptEmoji(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Trace___hyg_1552____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg(lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__8___boxed(lean_object**); +static lean_object* l_Lean_expandTraceMacro___closed__5; static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__14; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__2(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__50; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg___lambda__3(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__6___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__11; lean_object* lean_nat_mul(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_addTraceAsMessages___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__12(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__3; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1460____closed__2; lean_object* l___private_Init_Data_Array_QSort_0__Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace_profiler_threshold_unitAdjusted___boxed(lean_object*); +static lean_object* l_Lean_expandTraceMacro___closed__13; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1644____closed__1; static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368____closed__2; static double l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__1___closed__1; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_addTraceAsMessages___spec__8___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_BaseIO_toIO___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_addTraceAsMessages___spec__13(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resetTraceState___rarg___lambda__1___closed__1; @@ -538,19 +536,24 @@ static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__8; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_addTraceAsMessages___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_addTraceAsMessages___spec__15___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_exceptOptionEmoji___rarg(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__23; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption_go___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledForCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__17; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__11; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__7; static uint64_t l_Lean_instInhabitedTraceState___closed__1; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__13; LEAN_EXPORT lean_object* l_Lean_trace_profiler_output_pp; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__7___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessages___spec__11(lean_object*); +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_addTraceAsMessages___spec__16(lean_object*, size_t, size_t, lean_object*); @@ -563,22 +566,21 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTraceState___closed__5; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8; LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptMonadCacheT___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_modifyTraces___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_addTraceAsMessages___spec__13___boxed(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368____closed__3; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace_profiler_output; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__20; +static lean_object* l_Lean_expandTraceMacro___closed__8; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__28; LEAN_EXPORT lean_object* l_Lean_instExceptToEmojiBool(lean_object*); static double l_Lean_trace_profiler_threshold_unitAdjusted___closed__2; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__11; @@ -586,6 +588,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_addTraceAsMessag uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* l_instHashablePos___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptStateRefT_x27___rarg(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__49; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadTraceOfMonadLift___rarg(lean_object*, lean_object*); @@ -596,46 +599,45 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode(lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledForCore___closed__1; LEAN_EXPORT lean_object* l_Lean_setTraceState(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16; LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Util_Trace___hyg_3014____closed__24; LEAN_EXPORT lean_object* l_Lean_addTrace(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__33; uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_NameSSet_insert___spec__6(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42; lean_object* l_ReaderT_instMonadExceptOf___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_addTraceAsMessages___spec__3(lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1368_(lean_object*); static lean_object* l_Lean_resetTraceState___rarg___lambda__1___closed__4; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; lean_object* l_String_toSubstring_x27(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; +static lean_object* l_Lean_expandTraceMacro___closed__3; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__40; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__17; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__30; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12; LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_format(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__37; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Lean_instMonadAlwaysExceptReaderT(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_withTraceNode_x27___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__12; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; static lean_object* l_Lean_checkEmoji___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withStartStop___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); double lean_float_sub(double, double); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__21; LEAN_EXPORT lean_object* l_Lean_addTraceAsMessages___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; LEAN_EXPORT lean_object* l_Lean_modifyTraces(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; +static lean_object* l_Lean_expandTraceMacro___lambda__1___closed__1; lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_5____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_instInhabitedTraceElem___closed__1() { _start: @@ -6428,263 +6430,7 @@ x_6 = l_Lean_registerTraceClass(x_1, x_5, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("doElemTrace[_]__", 16, 16); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("andthen", 7, 7); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("trace[", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__5; -x_2 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ident", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__8; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__6; -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__9; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__11() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("]", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__11; -x_2 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__10; -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__12; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__14() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("orelse", 6, 6); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__14; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__16() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("interpolatedStr", 15, 15); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__16; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__18() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("term", 4, 4); -return x_1; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__18; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__20() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__19; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__21() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__17; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__20; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__22() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__15; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__21; -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__20; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__23() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; -x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__13; -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__22; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__24() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__2; -x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__23; -x_4 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_doElemTrace_x5b___x5d____() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__24; -return x_1; -} -} -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -6692,19 +6438,19 @@ x_1 = lean_mk_string_unchecked("doNested", 8, 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -6712,7 +6458,7 @@ x_1 = lean_mk_string_unchecked("do", 2, 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__4() { _start: { lean_object* x_1; @@ -6720,19 +6466,19 @@ x_1 = lean_mk_string_unchecked("doSeqIndent", 11, 11); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__6() { _start: { lean_object* x_1; @@ -6740,19 +6486,19 @@ x_1 = lean_mk_string_unchecked("doSeqItem", 9, 9); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__8() { _start: { lean_object* x_1; @@ -6760,19 +6506,19 @@ x_1 = lean_mk_string_unchecked("doLet", 5, 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__8; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__10() { _start: { lean_object* x_1; @@ -6780,7 +6526,7 @@ x_1 = lean_mk_string_unchecked("let", 3, 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__11() { _start: { lean_object* x_1; @@ -6788,19 +6534,19 @@ x_1 = lean_mk_string_unchecked("letDecl", 7, 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__13() { _start: { lean_object* x_1; @@ -6808,19 +6554,19 @@ x_1 = lean_mk_string_unchecked("letIdDecl", 9, 9); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__13; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -6828,26 +6574,26 @@ x_1 = lean_mk_string_unchecked("cls", 3, 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15; +x_1 = l_Lean_expandTraceMacro___lambda__1___closed__15; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__18() { _start: { lean_object* x_1; @@ -6855,7 +6601,7 @@ x_1 = lean_mk_string_unchecked(":=", 2, 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -6863,19 +6609,19 @@ x_1 = lean_mk_string_unchecked("doIf", 4, 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__19; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -6883,7 +6629,7 @@ x_1 = lean_mk_string_unchecked("if", 2, 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__22() { _start: { lean_object* x_1; @@ -6891,19 +6637,19 @@ x_1 = lean_mk_string_unchecked("doIfProp", 8, 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__22; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -6911,19 +6657,19 @@ x_1 = lean_mk_string_unchecked("paren", 5, 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__24; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -6931,7 +6677,7 @@ x_1 = lean_mk_string_unchecked("(", 1, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -6939,19 +6685,19 @@ x_1 = lean_mk_string_unchecked("liftMethod", 10, 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__27; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -6959,7 +6705,7 @@ x_1 = lean_mk_string_unchecked("←", 3, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__30() { _start: { lean_object* x_1; @@ -6967,19 +6713,19 @@ x_1 = lean_mk_string_unchecked("app", 3, 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__30; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__32() { _start: { lean_object* x_1; @@ -6987,16 +6733,16 @@ x_1 = lean_mk_string_unchecked("Lean.isTracingEnabledFor", 24, 24); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32; +x_1 = l_Lean_expandTraceMacro___lambda__1___closed__32; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__34() { _start: { lean_object* x_1; @@ -7004,41 +6750,41 @@ x_1 = lean_mk_string_unchecked("isTracingEnabledFor", 19, 19); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__34; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__35; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__36; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__38() { _start: { lean_object* x_1; @@ -7046,7 +6792,7 @@ x_1 = lean_mk_string_unchecked(")", 1, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__39() { _start: { lean_object* x_1; @@ -7054,7 +6800,7 @@ x_1 = lean_mk_string_unchecked("then", 4, 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__40() { _start: { lean_object* x_1; @@ -7062,19 +6808,19 @@ x_1 = lean_mk_string_unchecked("doExpr", 6, 6); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__40; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__42() { _start: { lean_object* x_1; @@ -7082,16 +6828,16 @@ x_1 = lean_mk_string_unchecked("Lean.addTrace", 13, 13); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42; +x_1 = l_Lean_expandTraceMacro___lambda__1___closed__42; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__44() { _start: { lean_object* x_1; @@ -7099,41 +6845,41 @@ x_1 = lean_mk_string_unchecked("addTrace", 8, 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__44; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__45; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; +x_2 = l_Lean_expandTraceMacro___lambda__1___closed__46; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__48() { _start: { lean_object* x_1; @@ -7141,19 +6887,19 @@ x_1 = lean_mk_string_unchecked("quotedName", 10, 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48; +x_4 = l_Lean_expandTraceMacro___lambda__1___closed__48; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__50() { _start: { lean_object* x_1; @@ -7161,7 +6907,7 @@ x_1 = lean_mk_string_unchecked(".", 1, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51() { +static lean_object* _init_l_Lean_expandTraceMacro___lambda__1___closed__51() { _start: { lean_object* x_1; @@ -7169,7 +6915,7 @@ x_1 = lean_mk_string_unchecked("`", 1, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; @@ -7183,12 +6929,12 @@ lean_inc(x_8); x_9 = lean_ctor_get(x_3, 1); lean_inc(x_9); lean_dec(x_3); -x_10 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; +x_10 = l_Lean_expandTraceMacro___lambda__1___closed__3; lean_inc(x_7); x_11 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_11, 0, x_7); lean_ctor_set(x_11, 1, x_10); -x_12 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10; +x_12 = l_Lean_expandTraceMacro___lambda__1___closed__10; lean_inc(x_7); x_13 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_13, 0, x_7); @@ -7200,19 +6946,19 @@ x_16 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_16, 0, x_7); lean_ctor_set(x_16, 1, x_14); lean_ctor_set(x_16, 2, x_15); -x_17 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17; +x_17 = l_Lean_expandTraceMacro___lambda__1___closed__17; lean_inc(x_8); lean_inc(x_9); x_18 = l_Lean_addMacroScope(x_9, x_17, x_8); x_19 = lean_box(0); -x_20 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16; +x_20 = l_Lean_expandTraceMacro___lambda__1___closed__16; lean_inc(x_7); x_21 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_21, 0, x_7); lean_ctor_set(x_21, 1, x_20); lean_ctor_set(x_21, 2, x_18); lean_ctor_set(x_21, 3, x_19); -x_22 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18; +x_22 = l_Lean_expandTraceMacro___lambda__1___closed__18; lean_inc(x_7); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_7); @@ -7221,27 +6967,27 @@ x_24 = l_Lean_Syntax_getId(x_1); x_25 = lean_erase_macro_scopes(x_24); lean_inc(x_25); x_26 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_25); -x_27 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; +x_27 = l_Lean_expandTraceMacro___lambda__1___closed__21; lean_inc(x_7); x_28 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_28, 0, x_7); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; +x_29 = l_Lean_expandTraceMacro___lambda__1___closed__26; lean_inc(x_7); x_30 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_30, 0, x_7); lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29; +x_31 = l_Lean_expandTraceMacro___lambda__1___closed__29; lean_inc(x_7); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_7); lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; +x_33 = l_Lean_expandTraceMacro___lambda__1___closed__35; lean_inc(x_8); lean_inc(x_9); x_34 = l_Lean_addMacroScope(x_9, x_33, x_8); -x_35 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33; -x_36 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37; +x_35 = l_Lean_expandTraceMacro___lambda__1___closed__33; +x_36 = l_Lean_expandTraceMacro___lambda__1___closed__37; lean_inc(x_7); x_37 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_37, 0, x_7); @@ -7251,33 +6997,33 @@ lean_ctor_set(x_37, 3, x_36); lean_inc(x_21); lean_inc(x_7); x_38 = l_Lean_Syntax_node1(x_7, x_14, x_21); -x_39 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31; +x_39 = l_Lean_expandTraceMacro___lambda__1___closed__31; lean_inc(x_7); x_40 = l_Lean_Syntax_node2(x_7, x_39, x_37, x_38); -x_41 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28; +x_41 = l_Lean_expandTraceMacro___lambda__1___closed__28; lean_inc(x_7); x_42 = l_Lean_Syntax_node2(x_7, x_41, x_32, x_40); -x_43 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; +x_43 = l_Lean_expandTraceMacro___lambda__1___closed__38; lean_inc(x_7); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_7); lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25; +x_45 = l_Lean_expandTraceMacro___lambda__1___closed__25; lean_inc(x_7); x_46 = l_Lean_Syntax_node3(x_7, x_45, x_30, x_42, x_44); -x_47 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23; +x_47 = l_Lean_expandTraceMacro___lambda__1___closed__23; lean_inc(x_16); lean_inc(x_7); x_48 = l_Lean_Syntax_node2(x_7, x_47, x_16, x_46); -x_49 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39; +x_49 = l_Lean_expandTraceMacro___lambda__1___closed__39; lean_inc(x_7); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_7); lean_ctor_set(x_50, 1, x_49); -x_51 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; +x_51 = l_Lean_expandTraceMacro___lambda__1___closed__45; x_52 = l_Lean_addMacroScope(x_9, x_51, x_8); -x_53 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43; -x_54 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47; +x_53 = l_Lean_expandTraceMacro___lambda__1___closed__43; +x_54 = l_Lean_expandTraceMacro___lambda__1___closed__47; lean_inc(x_7); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_7); @@ -7289,19 +7035,19 @@ lean_inc(x_7); x_56 = l_Lean_Syntax_node2(x_7, x_14, x_21, x_2); lean_inc(x_7); x_57 = l_Lean_Syntax_node2(x_7, x_39, x_55, x_56); -x_58 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41; +x_58 = l_Lean_expandTraceMacro___lambda__1___closed__41; lean_inc(x_7); x_59 = l_Lean_Syntax_node1(x_7, x_58, x_57); -x_60 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; +x_60 = l_Lean_expandTraceMacro___lambda__1___closed__7; lean_inc(x_16); lean_inc(x_7); x_61 = l_Lean_Syntax_node2(x_7, x_60, x_59, x_16); lean_inc(x_7); x_62 = l_Lean_Syntax_node1(x_7, x_14, x_61); -x_63 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5; +x_63 = l_Lean_expandTraceMacro___lambda__1___closed__5; lean_inc(x_7); x_64 = l_Lean_Syntax_node1(x_7, x_63, x_62); -x_65 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; +x_65 = l_Lean_expandTraceMacro___lambda__1___closed__20; lean_inc_n(x_16, 2); lean_inc(x_7); x_66 = l_Lean_Syntax_node6(x_7, x_65, x_28, x_48, x_50, x_64, x_16, x_16); @@ -7312,14 +7058,14 @@ if (lean_obj_tag(x_26) == 0) { lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; x_68 = l_Lean_quoteNameMk(x_25); -x_69 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; +x_69 = l_Lean_expandTraceMacro___lambda__1___closed__14; lean_inc_n(x_16, 2); lean_inc(x_7); x_70 = l_Lean_Syntax_node5(x_7, x_69, x_21, x_16, x_16, x_23, x_68); -x_71 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; +x_71 = l_Lean_expandTraceMacro___lambda__1___closed__12; lean_inc(x_7); x_72 = l_Lean_Syntax_node1(x_7, x_71, x_70); -x_73 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; +x_73 = l_Lean_expandTraceMacro___lambda__1___closed__9; lean_inc(x_16); lean_inc(x_7); x_74 = l_Lean_Syntax_node3(x_7, x_73, x_13, x_16, x_72); @@ -7329,7 +7075,7 @@ lean_inc(x_7); x_76 = l_Lean_Syntax_node2(x_7, x_14, x_75, x_67); lean_inc(x_7); x_77 = l_Lean_Syntax_node1(x_7, x_63, x_76); -x_78 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; +x_78 = l_Lean_expandTraceMacro___lambda__1___closed__2; x_79 = l_Lean_Syntax_node2(x_7, x_78, x_11, x_77); x_80 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_80, 0, x_79); @@ -7343,9 +7089,9 @@ lean_dec(x_25); x_81 = lean_ctor_get(x_26, 0); lean_inc(x_81); lean_dec(x_26); -x_82 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50; +x_82 = l_Lean_expandTraceMacro___lambda__1___closed__50; x_83 = l_String_intercalate(x_82, x_81); -x_84 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51; +x_84 = l_Lean_expandTraceMacro___lambda__1___closed__51; x_85 = lean_string_append(x_84, x_83); lean_dec(x_83); x_86 = lean_box(2); @@ -7354,19 +7100,19 @@ x_88 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_88, 0, x_87); lean_ctor_set(x_88, 1, x_19); x_89 = lean_array_mk(x_88); -x_90 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49; +x_90 = l_Lean_expandTraceMacro___lambda__1___closed__49; x_91 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_91, 0, x_86); lean_ctor_set(x_91, 1, x_90); lean_ctor_set(x_91, 2, x_89); -x_92 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; +x_92 = l_Lean_expandTraceMacro___lambda__1___closed__14; lean_inc_n(x_16, 2); lean_inc(x_7); x_93 = l_Lean_Syntax_node5(x_7, x_92, x_21, x_16, x_16, x_23, x_91); -x_94 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; +x_94 = l_Lean_expandTraceMacro___lambda__1___closed__12; lean_inc(x_7); x_95 = l_Lean_Syntax_node1(x_7, x_94, x_93); -x_96 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; +x_96 = l_Lean_expandTraceMacro___lambda__1___closed__9; lean_inc(x_16); lean_inc(x_7); x_97 = l_Lean_Syntax_node3(x_7, x_96, x_13, x_16, x_95); @@ -7376,7 +7122,7 @@ lean_inc(x_7); x_99 = l_Lean_Syntax_node2(x_7, x_14, x_98, x_67); lean_inc(x_7); x_100 = l_Lean_Syntax_node1(x_7, x_63, x_99); -x_101 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; +x_101 = l_Lean_expandTraceMacro___lambda__1___closed__2; x_102 = l_Lean_Syntax_node2(x_7, x_101, x_11, x_100); x_103 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_103, 0, x_102); @@ -7385,7 +7131,7 @@ return x_103; } } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__1() { _start: { lean_object* x_1; @@ -7393,17 +7139,17 @@ x_1 = lean_mk_string_unchecked("interpolatedStrKind", 19, 19); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1; +x_2 = l_Lean_expandTraceMacro___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__3() { _start: { lean_object* x_1; @@ -7411,19 +7157,19 @@ x_1 = lean_mk_string_unchecked("typeAscription", 14, 14); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; x_2 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__1; x_3 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__13; -x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3; +x_4 = l_Lean_expandTraceMacro___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__5() { _start: { lean_object* x_1; @@ -7431,7 +7177,7 @@ x_1 = lean_mk_string_unchecked(":", 1, 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__6() { _start: { lean_object* x_1; @@ -7439,82 +7185,82 @@ x_1 = lean_mk_string_unchecked("MessageData", 11, 11); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; +x_1 = l_Lean_expandTraceMacro___closed__6; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; +x_2 = l_Lean_expandTraceMacro___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; +x_2 = l_Lean_expandTraceMacro___closed__6; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; +x_2 = l_Lean_expandTraceMacro___closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; +x_1 = l_Lean_expandTraceMacro___closed__9; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11; +x_2 = l_Lean_expandTraceMacro___closed__11; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12; +x_1 = l_Lean_expandTraceMacro___closed__10; +x_2 = l_Lean_expandTraceMacro___closed__12; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__14() { _start: { lean_object* x_1; @@ -7522,17 +7268,17 @@ x_1 = lean_mk_string_unchecked("termM!_", 7, 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14; +x_2 = l_Lean_expandTraceMacro___closed__14; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16() { +static lean_object* _init_l_Lean_expandTraceMacro___closed__16() { _start: { lean_object* x_1; @@ -7540,6 +7286,354 @@ x_1 = lean_mk_string_unchecked("m!", 2, 2); return x_1; } } +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +lean_inc(x_2); +x_5 = l_Lean_Syntax_getKind(x_2); +x_6 = l_Lean_expandTraceMacro___closed__2; +x_7 = lean_name_eq(x_5, x_6); +lean_dec(x_5); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_8 = lean_ctor_get(x_3, 5); +lean_inc(x_8); +x_9 = 0; +x_10 = l_Lean_SourceInfo_fromRef(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_3, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +x_13 = l_Lean_expandTraceMacro___lambda__1___closed__26; +lean_inc(x_10); +x_14 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_14, 0, x_10); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Lean_expandTraceMacro___closed__5; +lean_inc(x_10); +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_expandTraceMacro___closed__8; +x_18 = l_Lean_addMacroScope(x_12, x_17, x_11); +x_19 = l_Lean_expandTraceMacro___closed__7; +x_20 = l_Lean_expandTraceMacro___closed__13; +lean_inc(x_10); +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_20); +x_22 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__8; +lean_inc(x_10); +x_23 = l_Lean_Syntax_node1(x_10, x_22, x_21); +x_24 = l_Lean_expandTraceMacro___lambda__1___closed__38; +lean_inc(x_10); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_10); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_expandTraceMacro___closed__4; +x_27 = l_Lean_Syntax_node5(x_10, x_26, x_14, x_2, x_16, x_23, x_25); +x_28 = l_Lean_expandTraceMacro___lambda__1(x_1, x_27, x_3, x_4); +return x_28; +} +else +{ +lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_29 = lean_ctor_get(x_3, 5); +lean_inc(x_29); +x_30 = 0; +x_31 = l_Lean_SourceInfo_fromRef(x_29, x_30); +lean_dec(x_29); +x_32 = l_Lean_expandTraceMacro___closed__16; +lean_inc(x_31); +x_33 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_expandTraceMacro___closed__15; +x_35 = l_Lean_Syntax_node2(x_31, x_34, x_33, x_2); +x_36 = l_Lean_expandTraceMacro___lambda__1(x_1, x_35, x_3, x_4); +return x_36; +} +} +} +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_expandTraceMacro___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_expandTraceMacro___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_expandTraceMacro(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("doElemTrace[_]__", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1276____closed__5; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("trace[", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__5; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ident", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__8; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__6; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__11; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__10; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__12; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("orelse", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("interpolatedStr", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__16; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__18; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__19; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__17; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__15; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__21; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__20; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__13; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__22; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__23; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d____() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__24; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -7560,143 +7654,33 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = lean_unsigned_to_nat(3u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); lean_dec(x_1); -lean_inc(x_11); -x_12 = l_Lean_Syntax_getKind(x_11); -x_13 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2; -x_14 = lean_name_eq(x_12, x_13); -lean_dec(x_12); -if (x_14 == 0) -{ -uint8_t x_54; -x_54 = 0; -x_15 = x_54; -goto block_53; -} -else -{ -uint8_t x_55; -x_55 = 1; -x_15 = x_55; -goto block_53; -} -block_53: -{ -if (x_15 == 0) -{ -lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_16 = lean_ctor_get(x_2, 5); -lean_inc(x_16); -x_17 = 0; -x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); -lean_dec(x_16); -x_19 = lean_ctor_get(x_2, 2); -lean_inc(x_19); -x_20 = lean_ctor_get(x_2, 1); -lean_inc(x_20); -x_21 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; -lean_inc(x_18); -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_18); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5; -lean_inc(x_18); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_18); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8; -x_26 = l_Lean_addMacroScope(x_20, x_25, x_19); -x_27 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7; -x_28 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13; -lean_inc(x_18); -x_29 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_29, 0, x_18); -lean_ctor_set(x_29, 1, x_27); -lean_ctor_set(x_29, 2, x_26); -lean_ctor_set(x_29, 3, x_28); -x_30 = l___auto____x40_Lean_Util_Trace___hyg_3014____closed__8; -lean_inc(x_18); -x_31 = l_Lean_Syntax_node1(x_18, x_30, x_29); -x_32 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; -lean_inc(x_18); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_18); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4; -x_35 = l_Lean_Syntax_node5(x_18, x_34, x_22, x_11, x_24, x_31, x_33); -x_36 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_9, x_35, x_2, x_3); +x_12 = l_Lean_expandTraceMacro(x_9, x_11, x_2, x_3); lean_dec(x_9); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) -{ -return x_36; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_36, 0); -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_36); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} -} -else -{ -lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_41 = lean_ctor_get(x_2, 5); -lean_inc(x_41); -x_42 = 0; -x_43 = l_Lean_SourceInfo_fromRef(x_41, x_42); -lean_dec(x_41); -x_44 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16; -lean_inc(x_43); -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15; -x_47 = l_Lean_Syntax_node2(x_43, x_46, x_45, x_11); -x_48 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_9, x_47, x_2, x_3); -lean_dec(x_9); -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -return x_48; +return x_12; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_48, 0); -x_51 = lean_ctor_get(x_48, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_48); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; -} -} -} -} +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_1, x_2, x_3, x_4); -lean_dec(x_1); -return x_5; } } static lean_object* _init_l_Lean_bombEmoji___closed__1() { @@ -11313,6 +11297,140 @@ l_Lean_registerTraceClass___closed__1 = _init_l_Lean_registerTraceClass___closed lean_mark_persistent(l_Lean_registerTraceClass___closed__1); l_Lean_registerTraceClass___closed__2 = _init_l_Lean_registerTraceClass___closed__2(); lean_mark_persistent(l_Lean_registerTraceClass___closed__2); +l_Lean_expandTraceMacro___lambda__1___closed__1 = _init_l_Lean_expandTraceMacro___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__1); +l_Lean_expandTraceMacro___lambda__1___closed__2 = _init_l_Lean_expandTraceMacro___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__2); +l_Lean_expandTraceMacro___lambda__1___closed__3 = _init_l_Lean_expandTraceMacro___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__3); +l_Lean_expandTraceMacro___lambda__1___closed__4 = _init_l_Lean_expandTraceMacro___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__4); +l_Lean_expandTraceMacro___lambda__1___closed__5 = _init_l_Lean_expandTraceMacro___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__5); +l_Lean_expandTraceMacro___lambda__1___closed__6 = _init_l_Lean_expandTraceMacro___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__6); +l_Lean_expandTraceMacro___lambda__1___closed__7 = _init_l_Lean_expandTraceMacro___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__7); +l_Lean_expandTraceMacro___lambda__1___closed__8 = _init_l_Lean_expandTraceMacro___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__8); +l_Lean_expandTraceMacro___lambda__1___closed__9 = _init_l_Lean_expandTraceMacro___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__9); +l_Lean_expandTraceMacro___lambda__1___closed__10 = _init_l_Lean_expandTraceMacro___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__10); +l_Lean_expandTraceMacro___lambda__1___closed__11 = _init_l_Lean_expandTraceMacro___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__11); +l_Lean_expandTraceMacro___lambda__1___closed__12 = _init_l_Lean_expandTraceMacro___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__12); +l_Lean_expandTraceMacro___lambda__1___closed__13 = _init_l_Lean_expandTraceMacro___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__13); +l_Lean_expandTraceMacro___lambda__1___closed__14 = _init_l_Lean_expandTraceMacro___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__14); +l_Lean_expandTraceMacro___lambda__1___closed__15 = _init_l_Lean_expandTraceMacro___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__15); +l_Lean_expandTraceMacro___lambda__1___closed__16 = _init_l_Lean_expandTraceMacro___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__16); +l_Lean_expandTraceMacro___lambda__1___closed__17 = _init_l_Lean_expandTraceMacro___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__17); +l_Lean_expandTraceMacro___lambda__1___closed__18 = _init_l_Lean_expandTraceMacro___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__18); +l_Lean_expandTraceMacro___lambda__1___closed__19 = _init_l_Lean_expandTraceMacro___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__19); +l_Lean_expandTraceMacro___lambda__1___closed__20 = _init_l_Lean_expandTraceMacro___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__20); +l_Lean_expandTraceMacro___lambda__1___closed__21 = _init_l_Lean_expandTraceMacro___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__21); +l_Lean_expandTraceMacro___lambda__1___closed__22 = _init_l_Lean_expandTraceMacro___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__22); +l_Lean_expandTraceMacro___lambda__1___closed__23 = _init_l_Lean_expandTraceMacro___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__23); +l_Lean_expandTraceMacro___lambda__1___closed__24 = _init_l_Lean_expandTraceMacro___lambda__1___closed__24(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__24); +l_Lean_expandTraceMacro___lambda__1___closed__25 = _init_l_Lean_expandTraceMacro___lambda__1___closed__25(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__25); +l_Lean_expandTraceMacro___lambda__1___closed__26 = _init_l_Lean_expandTraceMacro___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__26); +l_Lean_expandTraceMacro___lambda__1___closed__27 = _init_l_Lean_expandTraceMacro___lambda__1___closed__27(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__27); +l_Lean_expandTraceMacro___lambda__1___closed__28 = _init_l_Lean_expandTraceMacro___lambda__1___closed__28(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__28); +l_Lean_expandTraceMacro___lambda__1___closed__29 = _init_l_Lean_expandTraceMacro___lambda__1___closed__29(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__29); +l_Lean_expandTraceMacro___lambda__1___closed__30 = _init_l_Lean_expandTraceMacro___lambda__1___closed__30(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__30); +l_Lean_expandTraceMacro___lambda__1___closed__31 = _init_l_Lean_expandTraceMacro___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__31); +l_Lean_expandTraceMacro___lambda__1___closed__32 = _init_l_Lean_expandTraceMacro___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__32); +l_Lean_expandTraceMacro___lambda__1___closed__33 = _init_l_Lean_expandTraceMacro___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__33); +l_Lean_expandTraceMacro___lambda__1___closed__34 = _init_l_Lean_expandTraceMacro___lambda__1___closed__34(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__34); +l_Lean_expandTraceMacro___lambda__1___closed__35 = _init_l_Lean_expandTraceMacro___lambda__1___closed__35(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__35); +l_Lean_expandTraceMacro___lambda__1___closed__36 = _init_l_Lean_expandTraceMacro___lambda__1___closed__36(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__36); +l_Lean_expandTraceMacro___lambda__1___closed__37 = _init_l_Lean_expandTraceMacro___lambda__1___closed__37(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__37); +l_Lean_expandTraceMacro___lambda__1___closed__38 = _init_l_Lean_expandTraceMacro___lambda__1___closed__38(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__38); +l_Lean_expandTraceMacro___lambda__1___closed__39 = _init_l_Lean_expandTraceMacro___lambda__1___closed__39(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__39); +l_Lean_expandTraceMacro___lambda__1___closed__40 = _init_l_Lean_expandTraceMacro___lambda__1___closed__40(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__40); +l_Lean_expandTraceMacro___lambda__1___closed__41 = _init_l_Lean_expandTraceMacro___lambda__1___closed__41(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__41); +l_Lean_expandTraceMacro___lambda__1___closed__42 = _init_l_Lean_expandTraceMacro___lambda__1___closed__42(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__42); +l_Lean_expandTraceMacro___lambda__1___closed__43 = _init_l_Lean_expandTraceMacro___lambda__1___closed__43(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__43); +l_Lean_expandTraceMacro___lambda__1___closed__44 = _init_l_Lean_expandTraceMacro___lambda__1___closed__44(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__44); +l_Lean_expandTraceMacro___lambda__1___closed__45 = _init_l_Lean_expandTraceMacro___lambda__1___closed__45(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__45); +l_Lean_expandTraceMacro___lambda__1___closed__46 = _init_l_Lean_expandTraceMacro___lambda__1___closed__46(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__46); +l_Lean_expandTraceMacro___lambda__1___closed__47 = _init_l_Lean_expandTraceMacro___lambda__1___closed__47(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__47); +l_Lean_expandTraceMacro___lambda__1___closed__48 = _init_l_Lean_expandTraceMacro___lambda__1___closed__48(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__48); +l_Lean_expandTraceMacro___lambda__1___closed__49 = _init_l_Lean_expandTraceMacro___lambda__1___closed__49(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__49); +l_Lean_expandTraceMacro___lambda__1___closed__50 = _init_l_Lean_expandTraceMacro___lambda__1___closed__50(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__50); +l_Lean_expandTraceMacro___lambda__1___closed__51 = _init_l_Lean_expandTraceMacro___lambda__1___closed__51(); +lean_mark_persistent(l_Lean_expandTraceMacro___lambda__1___closed__51); +l_Lean_expandTraceMacro___closed__1 = _init_l_Lean_expandTraceMacro___closed__1(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__1); +l_Lean_expandTraceMacro___closed__2 = _init_l_Lean_expandTraceMacro___closed__2(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__2); +l_Lean_expandTraceMacro___closed__3 = _init_l_Lean_expandTraceMacro___closed__3(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__3); +l_Lean_expandTraceMacro___closed__4 = _init_l_Lean_expandTraceMacro___closed__4(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__4); +l_Lean_expandTraceMacro___closed__5 = _init_l_Lean_expandTraceMacro___closed__5(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__5); +l_Lean_expandTraceMacro___closed__6 = _init_l_Lean_expandTraceMacro___closed__6(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__6); +l_Lean_expandTraceMacro___closed__7 = _init_l_Lean_expandTraceMacro___closed__7(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__7); +l_Lean_expandTraceMacro___closed__8 = _init_l_Lean_expandTraceMacro___closed__8(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__8); +l_Lean_expandTraceMacro___closed__9 = _init_l_Lean_expandTraceMacro___closed__9(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__9); +l_Lean_expandTraceMacro___closed__10 = _init_l_Lean_expandTraceMacro___closed__10(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__10); +l_Lean_expandTraceMacro___closed__11 = _init_l_Lean_expandTraceMacro___closed__11(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__11); +l_Lean_expandTraceMacro___closed__12 = _init_l_Lean_expandTraceMacro___closed__12(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__12); +l_Lean_expandTraceMacro___closed__13 = _init_l_Lean_expandTraceMacro___closed__13(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__13); +l_Lean_expandTraceMacro___closed__14 = _init_l_Lean_expandTraceMacro___closed__14(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__14); +l_Lean_expandTraceMacro___closed__15 = _init_l_Lean_expandTraceMacro___closed__15(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__15); +l_Lean_expandTraceMacro___closed__16 = _init_l_Lean_expandTraceMacro___closed__16(); +lean_mark_persistent(l_Lean_expandTraceMacro___closed__16); l_Lean_doElemTrace_x5b___x5d_______closed__1 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__1(); lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__1); l_Lean_doElemTrace_x5b___x5d_______closed__2 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__2(); @@ -11363,140 +11481,6 @@ l_Lean_doElemTrace_x5b___x5d_______closed__24 = _init_l_Lean_doElemTrace_x5b___x lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__24); l_Lean_doElemTrace_x5b___x5d____ = _init_l_Lean_doElemTrace_x5b___x5d____(); lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d____); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__47); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__48); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__49); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__50); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16); l_Lean_bombEmoji___closed__1 = _init_l_Lean_bombEmoji___closed__1(); lean_mark_persistent(l_Lean_bombEmoji___closed__1); l_Lean_bombEmoji = _init_l_Lean_bombEmoji(); diff --git a/stage0/stdlib/Lean/Widget/InteractiveCode.c b/stage0/stdlib/Lean/Widget/InteractiveCode.c index c8f4c646881a..68e6ce6a0a58 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveCode.c +++ b/stage0/stdlib/Lean/Widget/InteractiveCode.c @@ -72,13 +72,13 @@ lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__3; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__29; LEAN_EXPORT lean_object* l_Lean_Widget_DiffTag_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_CodeWithInfos_mergePosMap___spec__3(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__21; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__15; extern lean_object* l_Lean_pp_raw; lean_object* l_Lean_SubExpr_Pos_toString(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____closed__1; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); @@ -99,7 +99,6 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonDiffTag; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552_(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_delab(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__22; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__5(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__26; @@ -156,6 +155,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_549_; lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableSubexprInfo_dec____x40_Lean_Widget_InteractiveCode___hyg_292____spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_pp_proofs; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_11____closed__8; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__1___closed__2; @@ -208,6 +208,7 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__5; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__1; lean_object* l_Lean_Widget_TaggedText_stripTags___rarg(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__4___closed__1; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__31; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____lambda__6___boxed(lean_object*, lean_object*, lean_object*); @@ -1410,6 +1411,32 @@ x_1 = l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_Interactiv return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +} +} LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552_(lean_object* x_1) { _start: { @@ -1435,7 +1462,7 @@ x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_7); x_12 = l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_353____closed__31; -x_13 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_12, x_4); +x_13 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_12, x_4); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_7); @@ -1446,11 +1473,20 @@ x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_8); lean_ctor_set(x_16, 1, x_15); x_17 = l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_68____closed__1; -x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_16, x_17); +x_18 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_16, x_17); x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____boxed(lean_object* x_1) { _start: { diff --git a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c index 7306f1166ffd..fd8a9afeea8f 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c +++ b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c @@ -41,11 +41,11 @@ static lean_object* l_Lean_Widget_instInhabitedMsgEmbed___closed__3; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_mkContextInfo___closed__1; lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____spec__2___closed__3; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_769____lambda__2___closed__3; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__2; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__45; -uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(lean_object*, lean_object*); static lean_object* l_Lean_Widget_InteractiveDiagnostic_toDiagnostic_prettyTt___closed__1; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__37; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_769____lambda__1(lean_object*, lean_object*, lean_object*); @@ -126,7 +126,6 @@ LEAN_EXPORT uint8_t l_Lean_Lsp_compareByUserVisible___at_Lean_Widget_Interactive lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg___closed__3; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__57; lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__3; @@ -135,6 +134,7 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget uint8_t lean_string_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractive_fmtToTT___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_msgToInteractive___closed__1; +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__13; static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg___closed__5; @@ -148,11 +148,11 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__54; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__34; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__2; -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg___closed__6; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_InteractiveDiagnostic___hyg_3275____closed__6; lean_object* lean_nat_to_int(lean_object*); lean_object* l_Subarray_size___rarg(lean_object*); +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instTypeNameLazyTraceChildren; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__6; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__51; @@ -198,7 +198,6 @@ static lean_object* l_Lean_Widget_InteractiveDiagnostic_toDiagnostic_prettyTt___ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__4(size_t, size_t, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__22; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Widget_instRpcEncodableMsgEmbed_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____lambda__1(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -285,6 +284,7 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_DiagnosticWith_ordUserVisible____x40_Lean_Data_Lsp_Diagnostics___hyg_2056____at_Lean_Widget_InteractiveDiagnostic_compareAsDiagnostics___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_chopUpChildren___closed__4; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_769____lambda__1___closed__1; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); @@ -367,6 +367,7 @@ lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1116_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_207_; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_769____lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractive_fmtToTT___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__55; @@ -377,9 +378,7 @@ uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_ordDiagnosticSeverity_ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__11; static lean_object* l_Lean_Widget_msgToInteractiveDiagnostic___closed__10; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__9; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__8; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__46; @@ -398,6 +397,7 @@ static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1853____rarg___closed__2; LEAN_EXPORT uint8_t l_Lean_Widget_InteractiveDiagnostic_compareAsDiagnostics(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___closed__8; @@ -428,10 +428,10 @@ static lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widge lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_infoview_maxTraceChildren; lean_object* l_Lean_Widget_TaggedText_stripTags___rarg(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__58; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableStrictOrLazy_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____rarg(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); lean_object* l_Lean_MessageData_format(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedStrictOrLazy___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_570____spec__5(size_t, size_t, lean_object*, lean_object*); @@ -2996,7 +2996,7 @@ if (x_8 == 0) lean_object* x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_array_size(x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_6, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_6, x_9); lean_ctor_set_tag(x_1, 4); lean_ctor_set(x_1, 0, x_11); x_12 = lean_alloc_ctor(0, 1, 0); @@ -3015,7 +3015,7 @@ lean_inc(x_15); lean_inc(x_14); lean_dec(x_7); x_16 = lean_array_size(x_14); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_16, x_6, x_14); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_16, x_6, x_14); lean_ctor_set_tag(x_1, 4); lean_ctor_set(x_1, 0, x_17); x_18 = lean_alloc_ctor(0, 1, 0); @@ -3050,7 +3050,7 @@ if (lean_is_exclusive(x_24)) { x_27 = lean_box(0); } x_28 = lean_array_size(x_25); -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_28, x_23, x_25); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_28, x_23, x_25); x_30 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_30, 0, x_29); x_31 = lean_alloc_ctor(0, 1, 0); @@ -4047,7 +4047,7 @@ lean_inc(x_10); lean_dec(x_9); x_11 = lean_array_size(x_10); x_12 = 0; -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_11, x_12, x_10); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_11, x_12, x_10); x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { @@ -4299,7 +4299,7 @@ lean_inc(x_64); lean_dec(x_63); x_65 = lean_array_size(x_64); x_66 = 0; -x_67 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_65, x_66, x_64); +x_67 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_65, x_66, x_64); x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); if (lean_is_exclusive(x_67)) { @@ -6283,16 +6283,16 @@ lean_ctor_set(x_6, 0, x_4); lean_ctor_set(x_6, 1, x_5); x_7 = lean_ctor_get(x_1, 1); x_8 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__23; -x_9 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_8, x_7); +x_9 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_8, x_7); x_10 = lean_ctor_get(x_1, 2); x_11 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__29; -x_12 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_11, x_10); +x_12 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_11, x_10); x_13 = lean_ctor_get(x_1, 3); x_14 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__35; -x_15 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_14, x_13); +x_15 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_14, x_13); x_16 = lean_ctor_get(x_1, 4); x_17 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__41; -x_18 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_17, x_16); +x_18 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_17, x_16); x_19 = lean_ctor_get(x_1, 5); x_20 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__47; lean_inc(x_19); @@ -6304,13 +6304,13 @@ lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_5); x_23 = lean_ctor_get(x_1, 6); x_24 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__52; -x_25 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_24, x_23); +x_25 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_24, x_23); x_26 = lean_ctor_get(x_1, 7); x_27 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__58; -x_28 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_27, x_26); +x_28 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_27, x_26); x_29 = lean_ctor_get(x_1, 8); x_30 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1938____closed__64; -x_31 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_30, x_29); +x_31 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_30, x_29); x_32 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_5); @@ -6339,7 +6339,7 @@ x_40 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_40, 0, x_6); lean_ctor_set(x_40, 1, x_39); x_41 = l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_2479____closed__1; -x_42 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_40, x_41); +x_42 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_40, x_41); x_43 = l_Lean_Json_mkObj(x_42); return x_43; } @@ -6566,7 +6566,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40 lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_165; x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_4); x_165 = lean_ctor_get(x_2, 1); lean_inc(x_165); if (lean_obj_tag(x_165) == 0) @@ -6585,7 +6585,7 @@ if (x_167 == 0) { lean_object* x_168; lean_object* x_169; x_168 = lean_ctor_get(x_165, 0); -x_169 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_168); +x_169 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_168); lean_ctor_set(x_165, 0, x_169); x_6 = x_165; x_7 = x_3; @@ -6597,7 +6597,7 @@ lean_object* x_170; lean_object* x_171; lean_object* x_172; x_170 = lean_ctor_get(x_165, 0); lean_inc(x_170); lean_dec(x_165); -x_171 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_170); +x_171 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_170); x_172 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_172, 0, x_171); x_6 = x_172; @@ -6897,7 +6897,7 @@ x_105 = lean_ctor_get(x_103, 1); lean_inc(x_105); lean_dec(x_103); x_106 = lean_array_size(x_104); -x_107 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_106, x_102, x_104); +x_107 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_106, x_102, x_104); x_108 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_108, 0, x_107); lean_ctor_set(x_97, 0, x_108); @@ -6920,7 +6920,7 @@ x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); lean_dec(x_112); x_115 = lean_array_size(x_113); -x_116 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_115, x_111, x_113); +x_116 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_115, x_111, x_113); x_117 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_117, 0, x_116); x_118 = lean_alloc_ctor(1, 1, 0); @@ -7043,7 +7043,7 @@ if (x_42 == 0) lean_object* x_43; size_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; x_43 = lean_ctor_get(x_41, 0); x_44 = lean_array_size(x_43); -x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_44, x_40, x_43); +x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_44, x_40, x_43); x_46 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_46, 0, x_45); lean_ctor_set(x_22, 0, x_46); @@ -7125,7 +7125,7 @@ lean_inc(x_59); lean_inc(x_58); lean_dec(x_41); x_60 = lean_array_size(x_58); -x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_60, x_40, x_58); +x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_60, x_40, x_58); x_62 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_62, 0, x_61); lean_ctor_set(x_22, 0, x_62); @@ -7212,7 +7212,7 @@ if (lean_is_exclusive(x_77)) { x_80 = lean_box(0); } x_81 = lean_array_size(x_78); -x_82 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_81, x_76, x_78); +x_82 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_81, x_76, x_78); x_83 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_83, 0, x_82); x_84 = lean_alloc_ctor(1, 1, 0); @@ -7865,7 +7865,7 @@ if (lean_is_exclusive(x_4)) { } x_7 = lean_ctor_get(x_5, 0); lean_inc(x_7); -x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_7); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_7); x_9 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_8, x_3); if (lean_obj_tag(x_9) == 0) { @@ -7919,7 +7919,7 @@ if (x_2049 == 0) { lean_object* x_2050; lean_object* x_2051; lean_object* x_2052; x_2050 = lean_ctor_get(x_2047, 0); -x_2051 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_2050); +x_2051 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_2050); x_2052 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_2051, x_3); if (lean_obj_tag(x_2052) == 0) { @@ -7964,7 +7964,7 @@ lean_object* x_2057; lean_object* x_2058; lean_object* x_2059; x_2057 = lean_ctor_get(x_2047, 0); lean_inc(x_2057); lean_dec(x_2047); -x_2058 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_2057); +x_2058 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_2057); x_2059 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_2058, x_3); if (lean_obj_tag(x_2059) == 0) { @@ -8538,7 +8538,7 @@ lean_inc(x_946); lean_dec(x_931); x_947 = lean_array_size(x_946); x_948 = 0; -x_949 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_947, x_948, x_946); +x_949 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_947, x_948, x_946); x_950 = lean_ctor_get(x_949, 0); lean_inc(x_950); lean_dec(x_949); @@ -8684,7 +8684,7 @@ lean_inc(x_985); lean_dec(x_970); x_986 = lean_array_size(x_985); x_987 = 0; -x_988 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_986, x_987, x_985); +x_988 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_986, x_987, x_985); x_989 = lean_ctor_get(x_988, 0); lean_inc(x_989); lean_dec(x_988); @@ -8974,7 +8974,7 @@ lean_inc(x_1008); lean_dec(x_1006); x_1009 = lean_array_size(x_1008); x_1010 = 0; -x_1011 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1009, x_1010, x_1008); +x_1011 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1009, x_1010, x_1008); x_1012 = !lean_is_exclusive(x_1011); if (x_1012 == 0) { @@ -9046,7 +9046,7 @@ x_1053 = lean_ctor_get(x_1052, 0); lean_inc(x_1053); lean_dec(x_1052); x_1054 = lean_array_size(x_1053); -x_1055 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1054, x_1010, x_1053); +x_1055 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1054, x_1010, x_1053); x_1056 = lean_ctor_get(x_1055, 0); lean_inc(x_1056); lean_dec(x_1055); @@ -9129,7 +9129,7 @@ x_1070 = lean_ctor_get(x_1069, 0); lean_inc(x_1070); lean_dec(x_1069); x_1071 = lean_array_size(x_1070); -x_1072 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1071, x_1010, x_1070); +x_1072 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1071, x_1010, x_1070); x_1073 = lean_ctor_get(x_1072, 0); lean_inc(x_1073); lean_dec(x_1072); @@ -9465,7 +9465,7 @@ x_1115 = lean_ctor_get(x_1113, 0); lean_inc(x_1115); lean_dec(x_1113); x_1116 = lean_array_size(x_1115); -x_1117 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1116, x_1010, x_1115); +x_1117 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1116, x_1010, x_1115); x_1118 = lean_ctor_get(x_1117, 0); lean_inc(x_1118); lean_dec(x_1117); @@ -9716,7 +9716,7 @@ lean_inc(x_1150); lean_dec(x_1148); x_1151 = lean_array_size(x_1150); x_1152 = 0; -x_1153 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1151, x_1152, x_1150); +x_1153 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1151, x_1152, x_1150); x_1154 = lean_ctor_get(x_1153, 0); lean_inc(x_1154); if (lean_is_exclusive(x_1153)) { @@ -9797,7 +9797,7 @@ x_1184 = lean_ctor_get(x_1182, 0); lean_inc(x_1184); lean_dec(x_1182); x_1185 = lean_array_size(x_1184); -x_1186 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1185, x_1152, x_1184); +x_1186 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1185, x_1152, x_1184); x_1187 = lean_ctor_get(x_1186, 0); lean_inc(x_1187); lean_dec(x_1186); @@ -10147,7 +10147,7 @@ lean_inc(x_1260); lean_dec(x_1259); x_1261 = lean_array_size(x_1260); x_1262 = 0; -x_1263 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1261, x_1262, x_1260); +x_1263 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1261, x_1262, x_1260); x_1264 = lean_ctor_get(x_1263, 0); lean_inc(x_1264); lean_dec(x_1263); @@ -10231,7 +10231,7 @@ lean_inc(x_1278); lean_dec(x_1277); x_1279 = lean_array_size(x_1278); x_1280 = 0; -x_1281 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1279, x_1280, x_1278); +x_1281 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1279, x_1280, x_1278); x_1282 = lean_ctor_get(x_1281, 0); lean_inc(x_1282); lean_dec(x_1281); @@ -10512,7 +10512,7 @@ lean_inc(x_1299); lean_dec(x_1297); x_1300 = lean_array_size(x_1299); x_1301 = 0; -x_1302 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1300, x_1301, x_1299); +x_1302 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1300, x_1301, x_1299); x_1303 = !lean_is_exclusive(x_1302); if (x_1303 == 0) { @@ -10585,7 +10585,7 @@ x_1344 = lean_ctor_get(x_1343, 0); lean_inc(x_1344); lean_dec(x_1343); x_1345 = lean_array_size(x_1344); -x_1346 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1345, x_1301, x_1344); +x_1346 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1345, x_1301, x_1344); x_1347 = lean_ctor_get(x_1346, 0); lean_inc(x_1347); lean_dec(x_1346); @@ -10670,7 +10670,7 @@ x_1361 = lean_ctor_get(x_1360, 0); lean_inc(x_1361); lean_dec(x_1360); x_1362 = lean_array_size(x_1361); -x_1363 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1362, x_1301, x_1361); +x_1363 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1362, x_1301, x_1361); x_1364 = lean_ctor_get(x_1363, 0); lean_inc(x_1364); lean_dec(x_1363); @@ -11011,7 +11011,7 @@ x_1406 = lean_ctor_get(x_1404, 0); lean_inc(x_1406); lean_dec(x_1404); x_1407 = lean_array_size(x_1406); -x_1408 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1407, x_1301, x_1406); +x_1408 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1407, x_1301, x_1406); x_1409 = lean_ctor_get(x_1408, 0); lean_inc(x_1409); lean_dec(x_1408); @@ -11266,7 +11266,7 @@ lean_inc(x_1441); lean_dec(x_1439); x_1442 = lean_array_size(x_1441); x_1443 = 0; -x_1444 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1442, x_1443, x_1441); +x_1444 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1442, x_1443, x_1441); x_1445 = lean_ctor_get(x_1444, 0); lean_inc(x_1445); if (lean_is_exclusive(x_1444)) { @@ -11348,7 +11348,7 @@ x_1475 = lean_ctor_get(x_1473, 0); lean_inc(x_1475); lean_dec(x_1473); x_1476 = lean_array_size(x_1475); -x_1477 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1476, x_1443, x_1475); +x_1477 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1476, x_1443, x_1475); x_1478 = lean_ctor_get(x_1477, 0); lean_inc(x_1478); lean_dec(x_1477); @@ -11666,7 +11666,7 @@ lean_inc(x_1531); lean_dec(x_1529); x_1532 = lean_array_size(x_1531); x_1533 = 0; -x_1534 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1532, x_1533, x_1531); +x_1534 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1532, x_1533, x_1531); x_1535 = lean_ctor_get(x_1534, 0); lean_inc(x_1535); lean_dec(x_1534); @@ -11878,7 +11878,7 @@ lean_inc(x_1553); lean_dec(x_1550); x_1554 = lean_array_size(x_1553); x_1555 = 0; -x_1556 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1554, x_1555, x_1553); +x_1556 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1554, x_1555, x_1553); x_1557 = lean_ctor_get(x_1556, 0); lean_inc(x_1557); if (lean_is_exclusive(x_1556)) { @@ -11965,7 +11965,7 @@ x_1587 = lean_ctor_get(x_1585, 0); lean_inc(x_1587); lean_dec(x_1585); x_1588 = lean_array_size(x_1587); -x_1589 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1588, x_1555, x_1587); +x_1589 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1588, x_1555, x_1587); x_1590 = lean_ctor_get(x_1589, 0); lean_inc(x_1590); lean_dec(x_1589); @@ -12329,7 +12329,7 @@ lean_inc(x_1651); lean_dec(x_1649); x_1652 = lean_array_size(x_1651); x_1653 = 0; -x_1654 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1652, x_1653, x_1651); +x_1654 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1652, x_1653, x_1651); x_1655 = lean_ctor_get(x_1654, 0); lean_inc(x_1655); lean_dec(x_1654); @@ -12547,7 +12547,7 @@ lean_inc(x_1673); lean_dec(x_1670); x_1674 = lean_array_size(x_1673); x_1675 = 0; -x_1676 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1674, x_1675, x_1673); +x_1676 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1674, x_1675, x_1673); x_1677 = lean_ctor_get(x_1676, 0); lean_inc(x_1677); if (lean_is_exclusive(x_1676)) { @@ -12634,7 +12634,7 @@ x_1707 = lean_ctor_get(x_1705, 0); lean_inc(x_1707); lean_dec(x_1705); x_1708 = lean_array_size(x_1707); -x_1709 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_1708, x_1675, x_1707); +x_1709 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_1708, x_1675, x_1707); x_1710 = lean_ctor_get(x_1709, 0); lean_inc(x_1710); lean_dec(x_1709); @@ -14148,7 +14148,7 @@ lean_inc(x_61); lean_dec(x_60); x_62 = lean_array_size(x_61); x_63 = 0; -x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_62, x_63, x_61); +x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_62, x_63, x_61); x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); lean_dec(x_64); @@ -14237,7 +14237,7 @@ lean_inc(x_80); lean_dec(x_79); x_81 = lean_array_size(x_80); x_82 = 0; -x_83 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_81, x_82, x_80); +x_83 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_81, x_82, x_80); x_84 = lean_ctor_get(x_83, 0); lean_inc(x_84); lean_dec(x_83); @@ -14523,7 +14523,7 @@ lean_inc(x_102); lean_dec(x_100); x_103 = lean_array_size(x_102); x_104 = 0; -x_105 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_103, x_104, x_102); +x_105 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_103, x_104, x_102); x_106 = !lean_is_exclusive(x_105); if (x_106 == 0) { @@ -14596,7 +14596,7 @@ x_147 = lean_ctor_get(x_146, 0); lean_inc(x_147); lean_dec(x_146); x_148 = lean_array_size(x_147); -x_149 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_148, x_104, x_147); +x_149 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_148, x_104, x_147); x_150 = lean_ctor_get(x_149, 0); lean_inc(x_150); lean_dec(x_149); @@ -14681,7 +14681,7 @@ x_164 = lean_ctor_get(x_163, 0); lean_inc(x_164); lean_dec(x_163); x_165 = lean_array_size(x_164); -x_166 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_165, x_104, x_164); +x_166 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_165, x_104, x_164); x_167 = lean_ctor_get(x_166, 0); lean_inc(x_167); lean_dec(x_166); @@ -15022,7 +15022,7 @@ x_209 = lean_ctor_get(x_207, 0); lean_inc(x_209); lean_dec(x_207); x_210 = lean_array_size(x_209); -x_211 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_210, x_104, x_209); +x_211 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_210, x_104, x_209); x_212 = lean_ctor_get(x_211, 0); lean_inc(x_212); lean_dec(x_211); @@ -15277,7 +15277,7 @@ lean_inc(x_244); lean_dec(x_242); x_245 = lean_array_size(x_244); x_246 = 0; -x_247 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_245, x_246, x_244); +x_247 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_245, x_246, x_244); x_248 = lean_ctor_get(x_247, 0); lean_inc(x_248); if (lean_is_exclusive(x_247)) { @@ -15359,7 +15359,7 @@ x_278 = lean_ctor_get(x_276, 0); lean_inc(x_278); lean_dec(x_276); x_279 = lean_array_size(x_278); -x_280 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_279, x_246, x_278); +x_280 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_279, x_246, x_278); x_281 = lean_ctor_get(x_280, 0); lean_inc(x_281); lean_dec(x_280); @@ -15717,7 +15717,7 @@ lean_inc(x_355); lean_dec(x_354); x_356 = lean_array_size(x_355); x_357 = 0; -x_358 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_356, x_357, x_355); +x_358 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_356, x_357, x_355); x_359 = lean_ctor_get(x_358, 0); lean_inc(x_359); lean_dec(x_358); @@ -15803,7 +15803,7 @@ lean_inc(x_373); lean_dec(x_372); x_374 = lean_array_size(x_373); x_375 = 0; -x_376 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_374, x_375, x_373); +x_376 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_374, x_375, x_373); x_377 = lean_ctor_get(x_376, 0); lean_inc(x_377); lean_dec(x_376); @@ -16088,7 +16088,7 @@ lean_inc(x_394); lean_dec(x_392); x_395 = lean_array_size(x_394); x_396 = 0; -x_397 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_395, x_396, x_394); +x_397 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_395, x_396, x_394); x_398 = !lean_is_exclusive(x_397); if (x_398 == 0) { @@ -16164,7 +16164,7 @@ x_441 = lean_ctor_get(x_440, 0); lean_inc(x_441); lean_dec(x_440); x_442 = lean_array_size(x_441); -x_443 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_442, x_396, x_441); +x_443 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_442, x_396, x_441); x_444 = lean_ctor_get(x_443, 0); lean_inc(x_444); lean_dec(x_443); @@ -16251,7 +16251,7 @@ x_458 = lean_ctor_get(x_457, 0); lean_inc(x_458); lean_dec(x_457); x_459 = lean_array_size(x_458); -x_460 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_459, x_396, x_458); +x_460 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_459, x_396, x_458); x_461 = lean_ctor_get(x_460, 0); lean_inc(x_461); lean_dec(x_460); @@ -16600,7 +16600,7 @@ x_505 = lean_ctor_get(x_503, 0); lean_inc(x_505); lean_dec(x_503); x_506 = lean_array_size(x_505); -x_507 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_506, x_396, x_505); +x_507 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_506, x_396, x_505); x_508 = lean_ctor_get(x_507, 0); lean_inc(x_508); lean_dec(x_507); @@ -16860,7 +16860,7 @@ lean_inc(x_540); lean_dec(x_538); x_541 = lean_array_size(x_540); x_542 = 0; -x_543 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_541, x_542, x_540); +x_543 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_541, x_542, x_540); x_544 = lean_ctor_get(x_543, 0); lean_inc(x_544); if (lean_is_exclusive(x_543)) { @@ -16945,7 +16945,7 @@ x_576 = lean_ctor_get(x_574, 0); lean_inc(x_576); lean_dec(x_574); x_577 = lean_array_size(x_576); -x_578 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_577, x_542, x_576); +x_578 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_577, x_542, x_576); x_579 = lean_ctor_get(x_578, 0); lean_inc(x_579); lean_dec(x_578); @@ -17270,7 +17270,7 @@ lean_inc(x_633); lean_dec(x_631); x_634 = lean_array_size(x_633); x_635 = 0; -x_636 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_634, x_635, x_633); +x_636 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_634, x_635, x_633); x_637 = lean_ctor_get(x_636, 0); lean_inc(x_637); lean_dec(x_636); @@ -17485,7 +17485,7 @@ lean_inc(x_655); lean_dec(x_652); x_656 = lean_array_size(x_655); x_657 = 0; -x_658 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_656, x_657, x_655); +x_658 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_656, x_657, x_655); x_659 = lean_ctor_get(x_658, 0); lean_inc(x_659); if (lean_is_exclusive(x_658)) { @@ -17575,7 +17575,7 @@ x_691 = lean_ctor_get(x_689, 0); lean_inc(x_691); lean_dec(x_689); x_692 = lean_array_size(x_691); -x_693 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_692, x_657, x_691); +x_693 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_692, x_657, x_691); x_694 = lean_ctor_get(x_693, 0); lean_inc(x_694); lean_dec(x_693); @@ -17947,7 +17947,7 @@ lean_inc(x_756); lean_dec(x_754); x_757 = lean_array_size(x_756); x_758 = 0; -x_759 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_757, x_758, x_756); +x_759 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_757, x_758, x_756); x_760 = lean_ctor_get(x_759, 0); lean_inc(x_760); lean_dec(x_759); @@ -18168,7 +18168,7 @@ lean_inc(x_778); lean_dec(x_775); x_779 = lean_array_size(x_778); x_780 = 0; -x_781 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_779, x_780, x_778); +x_781 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_779, x_780, x_778); x_782 = lean_ctor_get(x_781, 0); lean_inc(x_782); if (lean_is_exclusive(x_781)) { @@ -18258,7 +18258,7 @@ x_814 = lean_ctor_get(x_812, 0); lean_inc(x_814); lean_dec(x_812); x_815 = lean_array_size(x_814); -x_816 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_815, x_780, x_814); +x_816 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_815, x_780, x_814); x_817 = lean_ctor_get(x_816, 0); lean_inc(x_817); lean_dec(x_816); @@ -18851,7 +18851,7 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_2, 7); lean_inc(x_18); lean_dec(x_2); -x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_3, x_11); +x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_3, x_11); lean_dec(x_11); lean_dec(x_3); x_126 = lean_string_dec_lt(x_8, x_16); @@ -18925,7 +18925,7 @@ lean_dec(x_4); x_123 = lean_ctor_get(x_12, 0); lean_inc(x_123); lean_dec(x_12); -x_124 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_900_(x_122, x_123); +x_124 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordRange____x40_Lean_Data_Lsp_Basic___hyg_715_(x_122, x_123); lean_dec(x_123); lean_dec(x_122); x_21 = x_124; diff --git a/stage0/stdlib/Lean/Widget/InteractiveGoal.c b/stage0/stdlib/Lean/Widget/InteractiveGoal.c index 1e0b1177a983..43784d4de622 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveGoal.c +++ b/stage0/stdlib/Lean/Widget/InteractiveGoal.c @@ -37,6 +37,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveHypothesisBund lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_withGoalCtx___at_Lean_Widget_goalToInteractive___spec__8(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2301____closed__16; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_enc____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoals___closed__2; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__38; @@ -122,7 +123,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__27; LEAN_EXPORT lean_object* l_Lean_Widget_withGoalCtx___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__9; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__6; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2301____closed__4; @@ -133,6 +133,7 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ LEAN_EXPORT lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__30; static lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoals___closed__3; +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227_(lean_object*, lean_object*); static lean_object* l_Lean_Widget_goalToInteractive___lambda__1___closed__1; extern lean_object* l_Lean_Meta_pp_implementationDetailHyps; @@ -141,7 +142,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ uint8_t l_Lean_Expr_isSort(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__55; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle; -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Widget_goalToInteractive___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -193,7 +193,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__41; LEAN_EXPORT lean_object* l_Lean_Widget_instAppendInteractiveGoals; static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_657____closed__1; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__3; static lean_object* l_Lean_Widget_withGoalCtx___rarg___lambda__2___closed__2; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__2(lean_object*, lean_object*); @@ -273,6 +272,7 @@ LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodabl lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveGoals_enc____x40_Lean_Widget_InteractiveGoal___hyg_3343____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__17; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__34; static lean_object* l_Lean_Widget_InteractiveGoalCore_pretty___closed__2; static lean_object* l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__2; @@ -331,6 +331,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoal; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2301____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_InteractiveGoalCore_pretty(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2(size_t, size_t, lean_object*, lean_object*); @@ -339,8 +340,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2301____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Widget_goalToInteractive___spec__7(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__40; -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__10; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__60; @@ -356,6 +355,7 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__30; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__18; uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); lean_object* l_Lean_LocalContext_sanitizeNames(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__13; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__8; @@ -376,10 +376,10 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_goalToInt static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_3451____closed__1; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__51; lean_object* l_Lean_Widget_TaggedText_stripTags___rarg(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277_(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_goalToInteractive___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); extern lean_object* l_Lean_Widget_instTypeNameTermInfo; LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2611_; LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedInteractiveHypothesisBundle; @@ -1228,15 +1228,15 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_12); x_20 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__36; -x_21 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_20, x_5); +x_21 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_20, x_5); x_22 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__42; -x_23 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_22, x_6); +x_23 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_22, x_6); x_24 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__48; -x_25 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_24, x_7); +x_25 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_24, x_7); x_26 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__54; -x_27 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_26, x_8); +x_27 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_26, x_8); x_28 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__60; -x_29 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_28, x_9); +x_29 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_28, x_9); x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_12); @@ -1262,7 +1262,7 @@ x_37 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_37, 0, x_13); lean_ctor_set(x_37, 1, x_36); x_38 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_660____closed__1; -x_39 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_37, x_38); +x_39 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_37, x_38); x_40 = l_Lean_Json_mkObj(x_39); return x_40; } @@ -1601,7 +1601,7 @@ x_8 = lean_ctor_get(x_6, 1); lean_inc(x_8); lean_dec(x_6); x_9 = lean_array_size(x_7); -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_9, x_5, x_7); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_9, x_5, x_7); x_11 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_11, 0, x_10); x_12 = lean_ctor_get(x_1, 1); @@ -1614,7 +1614,7 @@ x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); x_17 = lean_array_size(x_15); -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_17, x_5, x_15); +x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_17, x_5, x_15); x_19 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_19, 0, x_18); x_20 = lean_ctor_get(x_1, 2); @@ -3165,7 +3165,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = lean_array_size(x_7); x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_8, x_9, x_7); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_8, x_9, x_7); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); @@ -3207,7 +3207,7 @@ x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); x_20 = lean_array_size(x_19); -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_20, x_9, x_19); +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_20, x_9, x_19); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); lean_dec(x_21); @@ -5592,7 +5592,7 @@ lean_inc(x_418); lean_dec(x_417); x_419 = lean_array_size(x_418); x_420 = 0; -x_421 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_419, x_420, x_418); +x_421 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_419, x_420, x_418); x_422 = lean_ctor_get(x_421, 0); lean_inc(x_422); lean_dec(x_421); @@ -5635,7 +5635,7 @@ x_430 = lean_ctor_get(x_428, 0); lean_inc(x_430); lean_dec(x_428); x_431 = lean_array_size(x_430); -x_432 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_431, x_420, x_430); +x_432 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_431, x_420, x_430); x_433 = lean_ctor_get(x_432, 0); lean_inc(x_433); lean_dec(x_432); @@ -7266,7 +7266,7 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_12); x_20 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__16; -x_21 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_20, x_5); +x_21 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_20, x_5); x_22 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1277____closed__22; lean_inc(x_6); x_23 = lean_alloc_ctor(0, 2, 0); @@ -7284,9 +7284,9 @@ x_27 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_12); x_28 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__54; -x_29 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_28, x_8); +x_29 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_28, x_8); x_30 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_176____closed__60; -x_31 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_30, x_9); +x_31 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_30, x_9); x_32 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_12); @@ -7312,7 +7312,7 @@ x_39 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_39, 0, x_13); lean_ctor_set(x_39, 1, x_38); x_40 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_660____closed__1; -x_41 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_39, x_40); +x_41 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_39, x_40); x_42 = l_Lean_Json_mkObj(x_41); return x_42; } @@ -7394,7 +7394,7 @@ x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); x_10 = lean_array_size(x_8); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_6, x_8); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_6, x_8); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = lean_ctor_get(x_3, 1); @@ -7904,7 +7904,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = lean_array_size(x_7); x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_8, x_9, x_7); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_8, x_9, x_7); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); @@ -9188,7 +9188,7 @@ lean_inc(x_250); lean_dec(x_249); x_251 = lean_array_size(x_250); x_252 = 0; -x_253 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_251, x_252, x_250); +x_253 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_251, x_252, x_250); x_254 = lean_ctor_get(x_253, 0); lean_inc(x_254); lean_dec(x_253); @@ -10270,7 +10270,7 @@ x_27 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_27, 0, x_10); lean_ctor_set(x_27, 1, x_26); x_28 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_660____closed__1; -x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_27, x_28); +x_29 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_27, x_28); x_30 = l_Lean_Json_mkObj(x_29); return x_30; } @@ -10317,7 +10317,7 @@ x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); x_10 = lean_array_size(x_8); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_10, x_6, x_8); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_10, x_6, x_8); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); x_13 = lean_ctor_get(x_3, 1); @@ -10342,7 +10342,7 @@ lean_inc(x_23); lean_dec(x_21); x_24 = lean_ctor_get(x_1, 1); lean_inc(x_24); -x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_24); +x_25 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_24); x_26 = lean_ctor_get(x_1, 2); lean_inc(x_26); lean_dec(x_1); @@ -10451,7 +10451,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = lean_array_size(x_7); x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_8, x_9, x_7); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_8, x_9, x_7); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); lean_dec(x_10); @@ -10583,7 +10583,7 @@ lean_inc(x_36); lean_dec(x_32); x_37 = lean_ctor_get(x_5, 3); lean_inc(x_37); -x_38 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_37); +x_38 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_37); x_39 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_38, x_2); if (lean_obj_tag(x_39) == 0) { @@ -10720,7 +10720,7 @@ lean_inc(x_66); lean_dec(x_65); x_67 = lean_array_size(x_66); x_68 = 0; -x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_67, x_68, x_66); +x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_67, x_68, x_66); x_70 = lean_ctor_get(x_69, 0); lean_inc(x_70); lean_dec(x_69); @@ -10856,7 +10856,7 @@ lean_inc(x_95); lean_dec(x_91); x_96 = lean_ctor_get(x_64, 3); lean_inc(x_96); -x_97 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_96); +x_97 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_96); x_98 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_97, x_2); if (lean_obj_tag(x_98) == 0) { @@ -12007,7 +12007,7 @@ x_6 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); x_7 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_660____closed__1; -x_8 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_6, x_7); +x_8 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221____spec__1(x_6, x_7); x_9 = l_Lean_Json_mkObj(x_8); return x_9; } @@ -12076,7 +12076,7 @@ if (x_6 == 0) lean_object* x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_7 = lean_ctor_get(x_5, 0); x_8 = lean_array_size(x_7); -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_8, x_4, x_7); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_8, x_4, x_7); x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); x_11 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_3454_(x_10); @@ -12092,7 +12092,7 @@ lean_inc(x_13); lean_inc(x_12); lean_dec(x_5); x_14 = lean_array_size(x_12); -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_14, x_4, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_14, x_4, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_3454_(x_16); @@ -12192,7 +12192,7 @@ lean_inc(x_6); lean_dec(x_5); x_7 = lean_array_size(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_7, x_8, x_6); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_7, x_8, x_6); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); lean_dec(x_9); @@ -12267,7 +12267,7 @@ lean_inc(x_26); lean_dec(x_25); x_27 = lean_array_size(x_26); x_28 = 0; -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_27, x_28, x_26); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_27, x_28, x_26); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); lean_dec(x_29); diff --git a/stage0/stdlib/Lean/Widget/UserWidget.c b/stage0/stdlib/Lean/Widget/UserWidget.c index 5d0c9e0cb0c4..027b8e68ea86 100644 --- a/stage0/stdlib/Lean/Widget/UserWidget.c +++ b/stage0/stdlib/Lean/Widget/UserWidget.c @@ -64,8 +64,10 @@ static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__18; lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__27; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); +lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__3; extern lean_object* l_Lean_maxRecDepthErrorMessage; @@ -147,7 +149,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Wid static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__5; static lean_object* l_Lean_Widget_getWidgetSource___lambda__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetGlobal___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* l_Lean_Environment_find_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Widget_showWidgetSpec___closed__2; lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1426____closed__2; @@ -223,7 +225,6 @@ extern lean_object* l_Lean_Lsp_instFromJsonPosition; static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4587____closed__1; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__32; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instToModuleModule; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -231,6 +232,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWi LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4953_(lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__7; size_t lean_usize_of_nat(lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3___closed__3; static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetGlobal(lean_object*); @@ -242,6 +244,7 @@ static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40 static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__19; +lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__5(uint64_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); lean_object* l_Lean_FileMap_lspPosToUtf8Pos(lean_object*, lean_object*); @@ -266,7 +269,6 @@ LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal___rarg(lean_object*, lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_List_toArray___rarg(lean_object*); -lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); @@ -346,7 +348,6 @@ lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Widget_addWidgetSpec___closed__7; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58; -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__12; static lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4842____closed__1; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2014____spec__2(lean_object*, lean_object*); @@ -489,12 +490,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJso lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instToModuleUserWidgetDefinition___boxed(lean_object*); static lean_object* l_Lean_Widget_widgetCmd___closed__4; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__4; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____closed__2; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__8; lean_object* l_Lean_Server_instRpcEncodableOfFromJsonOfToJson___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___lambda__1___closed__1; static lean_object* l_Lean_Widget_showWidgetSpec___closed__5; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__1; @@ -513,7 +514,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Wid uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_Environment_mainModule(lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3680____closed__1; lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalConstNoOverloadWithInfo___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_elabShowPanelWidgetsCmd___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -597,7 +598,6 @@ static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGe LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_getWidgets___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_showWidgetSpec; LEAN_EXPORT lean_object* l_Lean_Widget_widgetInfosAt_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); @@ -655,6 +655,7 @@ static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRp static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__6; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__1; static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__6; +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(size_t, size_t, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_moduleRegistry; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__9; @@ -677,9 +678,7 @@ static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240 LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_getWidgets___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__10; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__3; -lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__14; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_evalPanelWidgets(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -704,6 +703,7 @@ static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRp static lean_object* l_Lean_Widget_instToJsonWidgetSource___closed__1; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__28; static lean_object* l_Lean_Widget_eraseWidgetSpec___closed__1; @@ -717,6 +717,7 @@ static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36; LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4988_; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__3; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -750,7 +751,6 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonUserWidgetDefinition; lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1426____spec__1___boxed(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__2; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4842____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -8186,7 +8186,7 @@ lean_ctor_set(x_8, 1, x_7); x_9 = lean_ctor_get(x_1, 0); lean_inc(x_9); lean_dec(x_1); -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_9); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_221_(x_9); x_11 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_820____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -8201,7 +8201,7 @@ x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_8); lean_ctor_set(x_15, 1, x_14); x_16 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_15, x_16); +x_17 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_15, x_16); x_18 = l_Lean_Json_mkObj(x_17); return x_18; } @@ -8411,7 +8411,7 @@ x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_820____closed__2; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -8536,7 +8536,7 @@ x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_7, x_8); +x_9 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_7, x_8); x_10 = l_Lean_Json_mkObj(x_9); return x_10; } @@ -14039,7 +14039,8 @@ x_17 = lean_ctor_get(x_14, 1); x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); -x_19 = lean_environment_main_module(x_18); +x_19 = l_Lean_Environment_mainModule(x_18); +lean_dec(x_18); x_20 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__5; lean_inc(x_12); lean_ctor_set_tag(x_14, 2); @@ -14265,7 +14266,8 @@ lean_dec(x_14); x_118 = lean_ctor_get(x_116, 0); lean_inc(x_118); lean_dec(x_116); -x_119 = lean_environment_main_module(x_118); +x_119 = l_Lean_Environment_mainModule(x_118); +lean_dec(x_118); x_120 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__5; lean_inc(x_12); x_121 = lean_alloc_ctor(2, 2, 0); @@ -14716,7 +14718,8 @@ x_32 = lean_ctor_get(x_29, 1); x_33 = lean_ctor_get(x_31, 0); lean_inc(x_33); lean_dec(x_31); -x_34 = lean_environment_main_module(x_33); +x_34 = l_Lean_Environment_mainModule(x_33); +lean_dec(x_33); x_35 = l_Lean_Widget_elabWidgetInstanceSpec___closed__5; x_36 = l_Lean_addMacroScope(x_34, x_35, x_28); x_37 = l_Lean_Widget_elabWidgetInstanceSpec___closed__2; @@ -14765,7 +14768,8 @@ lean_dec(x_29); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_environment_main_module(x_54); +x_55 = l_Lean_Environment_mainModule(x_54); +lean_dec(x_54); x_56 = l_Lean_Widget_elabWidgetInstanceSpec___closed__5; x_57 = l_Lean_addMacroScope(x_55, x_56, x_28); x_58 = l_Lean_Widget_elabWidgetInstanceSpec___closed__2; @@ -15758,7 +15762,8 @@ x_28 = lean_ctor_get(x_25, 1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_12); +x_30 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_31 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); @@ -15940,7 +15945,8 @@ lean_dec(x_25); x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); lean_dec(x_77); -x_80 = lean_environment_main_module(x_12); +x_80 = l_Lean_Environment_mainModule(x_12); +lean_dec(x_12); x_81 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_81, 0, x_24); lean_ctor_set(x_81, 1, x_80); @@ -18086,7 +18092,8 @@ lean_dec(x_251); x_254 = lean_ctor_get(x_252, 0); lean_inc(x_254); lean_dec(x_252); -x_255 = lean_environment_main_module(x_254); +x_255 = l_Lean_Environment_mainModule(x_254); +lean_dec(x_254); x_256 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37; x_257 = l_Lean_addMacroScope(x_255, x_256, x_250); x_258 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36; @@ -18894,7 +18901,7 @@ x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_8); lean_ctor_set(x_14, 1, x_13); x_15 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_14, x_15); +x_16 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_14, x_15); x_17 = l_Lean_Json_mkObj(x_16); return x_17; } @@ -18932,7 +18939,7 @@ x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); x_31 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_30, x_31); +x_32 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_30, x_31); x_33 = l_Lean_Json_mkObj(x_32); return x_33; } @@ -19937,9 +19944,9 @@ x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_9); x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4203____closed__25; -x_18 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_17, x_5); +x_18 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_17, x_5); x_19 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3680____closed__1; -x_20 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_19, x_6); +x_20 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_552____spec__1(x_19, x_6); x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_9); @@ -19956,7 +19963,7 @@ x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_10); lean_ctor_set(x_25, 1, x_24); x_26 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_27 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_25, x_26); +x_27 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_25, x_26); x_28 = l_Lean_Json_mkObj(x_27); return x_28; } @@ -20155,7 +20162,7 @@ if (x_46 == 0) lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; x_47 = lean_ctor_get(x_13, 0); x_48 = lean_ctor_get(x_14, 0); -x_49 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_48); +x_49 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_48); lean_ctor_set(x_14, 0, x_49); x_50 = lean_ctor_get(x_1, 2); lean_inc(x_50); @@ -20227,7 +20234,7 @@ x_64 = lean_ctor_get(x_13, 0); x_65 = lean_ctor_get(x_14, 0); lean_inc(x_65); lean_dec(x_14); -x_66 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_65); +x_66 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_65); x_67 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_67, 0, x_66); x_68 = lean_ctor_get(x_1, 2); @@ -20298,7 +20305,7 @@ if (lean_is_exclusive(x_14)) { lean_dec_ref(x_14); x_81 = lean_box(0); } -x_82 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_80); +x_82 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_557_(x_80); if (lean_is_scalar(x_81)) { x_83 = lean_alloc_ctor(1, 1, 0); } else { @@ -20941,7 +20948,7 @@ if (x_62 == 0) { lean_object* x_63; lean_object* x_64; lean_object* x_65; x_63 = lean_ctor_get(x_60, 0); -x_64 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_63); +x_64 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_63); x_65 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_64, x_2); if (lean_obj_tag(x_65) == 0) { @@ -20985,7 +20992,7 @@ lean_object* x_70; lean_object* x_71; lean_object* x_72; x_70 = lean_ctor_get(x_60, 0); lean_inc(x_70); lean_dec(x_60); -x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_70); +x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_70); x_72 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_71, x_2); if (lean_obj_tag(x_72) == 0) { @@ -21271,7 +21278,7 @@ if (lean_is_exclusive(x_110)) { lean_dec_ref(x_110); x_113 = lean_box(0); } -x_114 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(x_112); +x_114 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_609_(x_112); x_115 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(x_114, x_2); if (lean_obj_tag(x_115) == 0) { @@ -21627,7 +21634,7 @@ x_6 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); x_7 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; -x_8 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_6, x_7); +x_8 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_CancelParams_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_CancelParams___hyg_86____spec__1(x_6, x_7); x_9 = l_Lean_Json_mkObj(x_8); return x_9; } @@ -21696,7 +21703,7 @@ if (x_6 == 0) lean_object* x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_7 = lean_ctor_get(x_5, 0); x_8 = lean_array_size(x_7); -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_8, x_4, x_7); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_8, x_4, x_7); x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); x_11 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4953_(x_10); @@ -21712,7 +21719,7 @@ lean_inc(x_13); lean_inc(x_12); lean_dec(x_5); x_14 = lean_array_size(x_12); -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_14, x_4, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1442____spec__2(x_14, x_4, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4953_(x_16); @@ -21817,7 +21824,7 @@ lean_inc(x_6); lean_dec(x_5); x_7 = lean_array_size(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_7, x_8, x_6); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_7, x_8, x_6); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); lean_dec(x_9); @@ -21891,7 +21898,7 @@ lean_inc(x_26); lean_dec(x_25); x_27 = lean_array_size(x_26); x_28 = 0; -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(x_27, x_28, x_26); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1498____spec__3(x_27, x_28, x_26); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); lean_dec(x_29); @@ -22117,9 +22124,8 @@ x_13 = lean_unsigned_to_nat(0u); x_14 = lean_array_uset(x_4, x_3, x_13); x_15 = lean_ctor_get(x_12, 0); lean_inc(x_15); -lean_inc(x_15); lean_inc(x_1); -x_16 = lean_environment_find(x_1, x_15); +x_16 = l_Lean_Environment_find_x3f(x_1, x_15); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; @@ -22340,9 +22346,8 @@ x_17 = lean_ctor_get(x_13, 1); lean_inc(x_17); x_45 = lean_ctor_get(x_16, 0); lean_inc(x_45); -lean_inc(x_45); lean_inc(x_2); -x_46 = lean_environment_find(x_2, x_45); +x_46 = l_Lean_Environment_find_x3f(x_2, x_45); if (lean_obj_tag(x_46) == 0) { lean_object* x_47; diff --git a/stage0/stdlib/Std/Sat/AIG/Basic.c b/stage0/stdlib/Std/Sat/AIG/Basic.c index 62ec80160257..b6c1749bf3be 100644 --- a/stage0/stdlib/Std/Sat/AIG/Basic.c +++ b/stage0/stdlib/Std/Sat/AIG/Basic.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std.Sat.AIG.Basic -// Imports: Std.Data.HashMap Std.Data.HashSet +// Imports: Std.Data.HashSet #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -6203,16 +6203,12 @@ lean_dec(x_1); return x_6; } } -lean_object* initialize_Std_Data_HashMap(uint8_t builtin, lean_object*); lean_object* initialize_Std_Data_HashSet(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Std_Sat_AIG_Basic(uint8_t builtin, lean_object* w) { lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Std_Data_HashMap(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Std_Data_HashSet(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.c b/stage0/stdlib/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.c index 12db8fc0a46b..78bcbf2b01c2 100644 --- a/stage0/stdlib/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.c +++ b/stage0/stdlib/Std/Tactic/BVDecide/LRAT/Internal/Formula/Implementation.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std.Tactic.BVDecide.LRAT.Internal.Formula.Implementation -// Imports: Init.Data.Array Std.Tactic.BVDecide.LRAT.Internal.Formula.Class Std.Tactic.BVDecide.LRAT.Internal.Assignment Std.Sat.CNF.Basic +// Imports: Std.Tactic.BVDecide.LRAT.Internal.Formula.Class Std.Tactic.BVDecide.LRAT.Internal.Assignment Std.Sat.CNF.Basic #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -5748,7 +5748,6 @@ lean_dec(x_1); return x_3; } } -lean_object* initialize_Init_Data_Array(uint8_t builtin, lean_object*); lean_object* initialize_Std_Tactic_BVDecide_LRAT_Internal_Formula_Class(uint8_t builtin, lean_object*); lean_object* initialize_Std_Tactic_BVDecide_LRAT_Internal_Assignment(uint8_t builtin, lean_object*); lean_object* initialize_Std_Sat_CNF_Basic(uint8_t builtin, lean_object*); @@ -5757,9 +5756,6 @@ LEAN_EXPORT lean_object* initialize_Std_Tactic_BVDecide_LRAT_Internal_Formula_Im lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Init_Data_Array(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Std_Tactic_BVDecide_LRAT_Internal_Formula_Class(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Std/Time/DateTime/Timestamp.c b/stage0/stdlib/Std/Time/DateTime/Timestamp.c index ea8611e7c670..200a12f05bc6 100644 --- a/stage0/stdlib/Std/Time/DateTime/Timestamp.c +++ b/stage0/stdlib/Std/Time/DateTime/Timestamp.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std.Time.DateTime.Timestamp -// Imports: Std.Time.Internal Init.Data.Int Init.System.IO Std.Time.Time Std.Time.Date Std.Time.Duration +// Imports: Std.Time.Internal Init.System.IO Std.Time.Time Std.Time.Date Std.Time.Duration #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -2101,7 +2101,6 @@ return x_3; } } lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); -lean_object* initialize_Init_Data_Int(uint8_t builtin, lean_object*); lean_object* initialize_Init_System_IO(uint8_t builtin, lean_object*); lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); @@ -2114,9 +2113,6 @@ _G_initialized = true; res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Init_Data_Int(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Init_System_IO(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Std/Time/Internal/Bounded.c b/stage0/stdlib/Std/Time/Internal/Bounded.c index 0f71ba77fdca..b2b3b98e2994 100644 --- a/stage0/stdlib/Std/Time/Internal/Bounded.c +++ b/stage0/stdlib/Std/Time/Internal/Bounded.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std.Time.Internal.Bounded -// Imports: Init.Data.Int +// Imports: Init.Omega #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -1938,13 +1938,13 @@ lean_dec(x_1); return x_3; } } -lean_object* initialize_Init_Data_Int(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Omega(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Std_Time_Internal_Bounded(uint8_t builtin, lean_object* w) { lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Init_Data_Int(builtin, lean_io_mk_world()); +res = initialize_Init_Omega(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1 = _init_l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1();